{-# OPTIONS --safe #-}

open import Ledger.Conway.Types.GovStructure
open import Ledger.Conway.Transaction using (TransactionStructure)

module Ledger.Conway.Gov (txs : _) (open TransactionStructure txs hiding (epoch)) where

open import Ledger.Prelude hiding (any?; Any; all?; All; Rel; lookup; ∈-filter)

open import Axiom.Set.Properties th using (∃-sublist-⇔)

open import Ledger.Conway.GovernanceActions govStructure using (Vote)
open import Ledger.Conway.Enact govStructure
open import Ledger.Conway.Ratify txs hiding (vote)
open import Ledger.Conway.Certs govStructure

open import stdlib.Data.List.Subpermutations using (subpermutations; sublists)
open import stdlib.Data.List.Subpermutations.Properties
open import Data.List.Membership.Propositional.Properties using (Any↔; ∈-filter⁻; ∈-filter⁺)
open import Data.List.Relation.Binary.Subset.Propositional using () renaming (_⊆_ to _⊆ˡ_)
open import Data.List.Relation.Unary.All using (all?; All)
open import Data.List.Relation.Unary.Any using (any?; Any)
open import Data.List.Relation.Unary.Unique.DecPropositional using (unique?)
open import Data.List.Relation.Unary.Unique.Propositional using (Unique)
open import Relation.Nullary.Decidable using () renaming (map to map-Dec)
open import Function.Properties.Equivalence using () renaming (sym to sym-Equiv)
open import Function.Related.Propositional using (↔⇒)

open GovActionState


GovState : Type


GovState = List (GovActionID × GovActionState)


record HasGovState {a} (A : Type a) : Type a where
  field GovStateOf : A  GovState
open HasGovState ⦃...⦄ public



record GovEnv : Type where
  field
    txid        : TxId
    epoch       : Epoch
    pparams     : PParams
    ppolicy     : Maybe ScriptHash
    enactState  : EnactState
    certState   : CertState
    rewardCreds :  Credential


instance
  unquoteDecl HasCast-GovEnv = derive-HasCast
    [ (quote GovEnv , HasCast-GovEnv) ]

private variable
  Γ : GovEnv
  s s' : GovState
  aid : GovActionID
  voter : Voter
  vote : GovVote
  v : Vote
  d : Coin
  addr : RwdAddr
  a : GovAction
  prev : NeedsHash (gaType a)
  k : 
  p : Maybe ScriptHash

open GState
open PState


govActionPriority : GovActionType  
govActionPriority NoConfidence     = 0
govActionPriority UpdateCommittee  = 1
govActionPriority NewConstitution  = 2
govActionPriority TriggerHF        = 3
govActionPriority ChangePParams    = 4
govActionPriority TreasuryWdrl     = 5
govActionPriority Info             = 6

Overlap : GovActionType  GovActionType  Type
Overlap NoConfidence    UpdateCommittee  = 
Overlap UpdateCommittee NoConfidence     = 
Overlap a               a'               = a  a'


-- TODO: cleanup this
Overlap? : (a a' : GovActionType)  Dec (Overlap a a')
Overlap? NoConfidence    UpdateCommittee  = Dec-⊤ .dec
Overlap? UpdateCommittee NoConfidence     = Dec-⊤ .dec
Overlap? NoConfidence NoConfidence = yes refl
Overlap? NoConfidence NewConstitution = no  ())
Overlap? NoConfidence TriggerHF = no  ())
Overlap? NoConfidence ChangePParams = no  ())
Overlap? NoConfidence TreasuryWdrl = no  ())
Overlap? NoConfidence Info = no  ())
Overlap? UpdateCommittee UpdateCommittee = yes refl
Overlap? UpdateCommittee NewConstitution = no  ())
Overlap? UpdateCommittee TriggerHF = no  ())
Overlap? UpdateCommittee ChangePParams = no  ())
Overlap? UpdateCommittee TreasuryWdrl = no  ())
Overlap? UpdateCommittee Info = no  ())
Overlap? NewConstitution NoConfidence = no  ())
Overlap? NewConstitution UpdateCommittee = no  ())
Overlap? NewConstitution NewConstitution = yes refl
Overlap? NewConstitution TriggerHF = no  ())
Overlap? NewConstitution ChangePParams = no  ())
Overlap? NewConstitution TreasuryWdrl = no  ())
Overlap? NewConstitution Info = no  ())
Overlap? TriggerHF NoConfidence = no  ())
Overlap? TriggerHF UpdateCommittee = no  ())
Overlap? TriggerHF NewConstitution = no  ())
Overlap? TriggerHF TriggerHF = yes refl
Overlap? TriggerHF ChangePParams = no  ())
Overlap? TriggerHF TreasuryWdrl = no  ())
Overlap? TriggerHF Info = no  ())
Overlap? ChangePParams NoConfidence = no  ())
Overlap? ChangePParams UpdateCommittee = no  ())
Overlap? ChangePParams NewConstitution = no  ())
Overlap? ChangePParams TriggerHF = no  ())
Overlap? ChangePParams ChangePParams = yes refl
Overlap? ChangePParams TreasuryWdrl = no  ())
Overlap? ChangePParams Info = no  ())
Overlap? TreasuryWdrl NoConfidence = no  ())
Overlap? TreasuryWdrl UpdateCommittee = no  ())
Overlap? TreasuryWdrl NewConstitution = no  ())
Overlap? TreasuryWdrl TriggerHF = no  ())
Overlap? TreasuryWdrl ChangePParams = no  ())
Overlap? TreasuryWdrl TreasuryWdrl = yes refl
Overlap? TreasuryWdrl Info = no  ())
Overlap? Info NoConfidence = no  ())
Overlap? Info UpdateCommittee = no  ())
Overlap? Info NewConstitution = no  ())
Overlap? Info TriggerHF = no  ())
Overlap? Info ChangePParams = no  ())
Overlap? Info TreasuryWdrl = no  ())
Overlap? Info Info = yes refl



insertGovAction : GovState  GovActionID × GovActionState  GovState
insertGovAction [] gaPr = [ gaPr ]
insertGovAction ((gaID₀ , gaSt₀)  gaPrs) (gaID₁ , gaSt₁)
  =  if govActionPriority (action gaSt₀ .gaType)  govActionPriority (action gaSt₁ .gaType)
     then (gaID₀ , gaSt₀)  insertGovAction gaPrs (gaID₁ , gaSt₁)
     else (gaID₁ , gaSt₁)  (gaID₀ , gaSt₀)  gaPrs

mkGovStatePair : Epoch  GovActionID  RwdAddr  (a : GovAction)  NeedsHash (a .gaType)
                  GovActionID × GovActionState
mkGovStatePair e aid addr a prev = (aid , record
  { votes =  ; returnAddr = addr ; expiresIn = e ; action = a ; prevAction = prev })

addAction : GovState
           Epoch  GovActionID  RwdAddr  (a : GovAction)  NeedsHash (a .gaType)
           GovState
addAction s e aid addr a prev = insertGovAction s (mkGovStatePair e aid addr a prev)


opaque


  addVote : GovState  GovActionID  Voter  Vote  GovState
  addVote s aid voter v = map modifyVotes s
    where modifyVotes : GovActionID × GovActionState  GovActionID × GovActionState
          modifyVotes = λ (gid , s')  gid , record s'
            { votes = if gid  aid then insert (votes s') voter v else votes s'}

  isRegistered : GovEnv  Voter  Type
  isRegistered Γ (r , c) = case r of


    λ where


      CC     just c  range (gState .ccHotKeys)
      DRep   c  dom (gState .dreps)
      SPO    c  mapˢ KeyHashObj (dom (pState .pools))
        where
          open CertState (GovEnv.certState Γ) using (gState; pState)

  validHFAction : GovProposal  GovState  EnactState  Type
  validHFAction (record { action = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#863}{\htmlId{6597}{\htmlClass{InductiveConstructor}{\text{TriggerHF}}}}\, \\ \,\href{Ledger.Conway.Gov.html#6609}{\htmlId{6609}{\htmlClass{Bound}{\text{v}}}}\, \end{pmatrix}$ ; prevAction = prev }) s e =
    (let (v' , aid) = EnactState.pv e in aid  prev × pvCanFollow v' v)
     ∃₂[ x , v' ] (prev , x)  fromList s × x .action  $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#863}{\htmlId{6775}{\htmlClass{InductiveConstructor}{\text{TriggerHF}}}}\, \\ \,\href{Ledger.Conway.Gov.html#6730}{\htmlId{6787}{\htmlClass{Bound}{\text{v'}}}}\, \end{pmatrix}$ × pvCanFollow v' v
  validHFAction _ _ _ = 


data


  _⊢_⇀⦇_,GOV⦈_  : GovEnv ×   GovState  GovVote  GovProposal  GovState  Type


_⊢_⇀⦇_,GOVS⦈_   : GovEnv  GovState  List (GovVote  GovProposal)  GovState  Type


-- Convert list of (GovActionID,GovActionState)-pairs to list of GovActionID pairs.
getAidPairsList : GovState  List (GovActionID × GovActionID)
getAidPairsList aid×states =
  mapMaybe  (aid , aState)  (aid ,_) <$> getHash (prevAction aState)) $ aid×states

-- A list of GovActionID pairs connects the first GovActionID to the second.
_connects_to_ : List (GovActionID × GovActionID)  GovActionID  GovActionID  Type
[] connects aidNew to aidOld = aidNew  aidOld
((aid , aidPrev)  s) connects aidNew to aidOld  =
  aid  aidNew × s connects aidPrev to aidOld  s connects aidNew to aidOld


enactable  : EnactState  List (GovActionID × GovActionID)
            GovActionID × GovActionState  Type
enactable e aidPairs = λ (aidNew , as)  case getHashES e (action as .gaType) of


  λ where


   nothing         
   (just aidOld)   ∃[ t ]  fromList t  fromList aidPairs
                            × Unique t × t connects aidNew to aidOld

allEnactable : EnactState  GovState  Type
allEnactable e aid×states = All (enactable e (getAidPairsList aid×states)) aid×states

hasParentE : EnactState  GovActionID  GovActionType  Type
hasParentE e aid gaTy = case getHashES e gaTy of


  λ where


   nothing     
   (just id)   id  aid

hasParent : EnactState  GovState  (gaTy : GovActionType)  NeedsHash gaTy  Type
hasParent e s gaTy aid = case getHash aid of


  λ where


    nothing       
    (just aid')   hasParentE e aid' gaTy
                    Any  (gid , gas)  gid  aid' × Overlap (gas .action .gaType) gaTy) s


open Equivalence

hasParentE? :  e aid a  Dec (hasParentE e aid a)
hasParentE? e aid gaTy with getHashES e gaTy
... | nothing   = yes _
... | (just id) = id  aid

hasParent? :  e s a aid  Dec (hasParent e s a aid)
hasParent? e s gaTy aid with getHash aid
... | just aid' = hasParentE? e aid' gaTy
                  ⊎-dec any?  (gid , gas)  gid  aid' ×-dec Overlap? (gas .action .gaType) gaTy) s
... | nothing = yes _

-- newtype to make the instance resolution work
data hasParent' : EnactState  GovState  (gaTy : GovActionType)  NeedsHash gaTy  Type where
  HasParent' :  {x y z w}  hasParent x y z w  hasParent' x y z w

instance
  hasParent?' :  {x y z w}  hasParent' x y z w 
  hasParent?' =  map′ HasParent'  where (HasParent' x)  x) (hasParent? _ _ _ _)

[_connects_to_?] :  l aidNew aidOld  Dec (l connects aidNew to aidOld)
[ [] connects aidNew to aidOld ?] = aidNew  aidOld

[ (aid , aidPrev)  s connects aidNew to aidOld ?] =
  ((aid  aidNew) ×-dec [ s connects aidPrev to aidOld ?]) ⊎-dec [ s connects aidNew to aidOld ?]

any?-connecting-subperm :  {u} {v}   L  Dec (Any l  Unique l × l connects u to v) (subpermutations L))
any?-connecting-subperm {u} {v} L = any?  l  unique? _≟_ l ×-dec [ l connects u to v ?]) (subpermutations L)

∃?-connecting-subperm :  {u} {v}   L  Dec (∃[ l ] l ∈ˡ subpermutations L × Unique l × l connects u to v)
∃?-connecting-subperm L = map-Dec (sym-Equiv (↔⇒ Any↔)) (any?-connecting-subperm L)

∃?-connecting-subset :  {u} {v}   L  Dec (∃[ l ] l ⊆ˡ L × Unique l × l connects u to v)
∃?-connecting-subset L = map-Dec (sym-Equiv ∃uniqueSubset⇔∃uniqueSubperm) (∃?-connecting-subperm L)

enactable? :  eState aidPairs aidNew×st  Dec (enactable eState aidPairs aidNew×st)
enactable? eState aidPairs (aidNew , as) with getHashES eState (GovActionState.action as .gaType)
... | nothing = yes tt
... | just aidOld = map-Dec (sym-Equiv ∃-sublist-⇔) (∃?-connecting-subset aidPairs)

allEnactable? :  eState aid×states  Dec (allEnactable eState aid×states)
allEnactable? eState aid×states =
  all?  aid×st  enactable? eState (getAidPairsList aid×states) aid×st) aid×states

-- newtype to make the instance resolution work
data allEnactable' : EnactState  GovState  Type where
  AllEnactable' :  {x y}  allEnactable x y  allEnactable' x y

instance
  allEnactable?' :  {x y}  allEnactable' x y 
  allEnactable?' =  map′ AllEnactable'  where (AllEnactable' x)  x) (allEnactable? _ _)

-- `maxAllEnactable` returns a list `ls` of sublists of the given
-- list (`aid×states : List (GovActionID × GovActionState)`) such that
--    (i) each sublist `l ∈ ls` satisfies `allEnactable e l` and
--   (ii) each sublist `l ∈ ls` is of maximal length among sublists of `aid×states` satisfying `allEnactable`.
maxAllEnactable : EnactState  List (GovActionID × GovActionState)  List (List (GovActionID × GovActionState))
maxAllEnactable e = maxsublists⊧P (allEnactable? e)

-- Every sublist returned by `maxAllEnactable` satisfies (i).
∈-maxAllEnactable→allEnactable :  {e} {aid×states} l
   l ∈ˡ maxAllEnactable e aid×states  allEnactable e l
∈-maxAllEnactable→allEnactable {e} {aid×states} l l∈ =
  proj₂ (∈-filter⁻ (allEnactable? e) {l} {sublists aid×states}
          (proj₁ (∈-filter⁻  l  length l  maxlen (sublists⊧P (allEnactable? e) aid×states)) l∈)))

-- Every sublist returned by `maxAllEnactable` satisfies (ii).
∈-maxAllEnactable→maxLength :  {e aid×states l l'}
                               l ∈ˡ sublists aid×states  allEnactable e l
                               l' ∈ˡ maxAllEnactable e aid×states
                               length l  length l'
∈-maxAllEnactable→maxLength {e} {aid×states} {l} {l'} l∈ el l'∈ =
  let ls = sublists⊧P (allEnactable? e) aid×states in
    subst (length l ≤_)
          (sym (proj₂ (∈-filter⁻  l  length l  maxlen ls) {xs = ls} l'∈)))
          (∈-maxlen-≤ l (∈-filter⁺ (allEnactable? e) l∈ el))


actionValid :  Credential  Maybe ScriptHash  Maybe ScriptHash  Epoch  GovAction  Type
actionValid rewardCreds p ppolicy epoch $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#898}{\htmlId{12653}{\htmlClass{InductiveConstructor}{\text{ChangePParams}}}}\, \\ \,\htmlId{12669}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$ =
  p  ppolicy
actionValid rewardCreds p ppolicy epoch $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#933}{\htmlId{12733}{\htmlClass{InductiveConstructor}{\text{TreasuryWdrl}}}}\,  \\ \,\href{Ledger.Conway.Gov.html#12749}{\htmlId{12749}{\htmlClass{Bound}{\text{x}}}}\, \end{pmatrix}$ =
  p  ppolicy × mapˢ RwdAddr.stake (dom x)  rewardCreds
actionValid rewardCreds p ppolicy epoch $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#793}{\htmlId{12856}{\htmlClass{InductiveConstructor}{\text{UpdateCommittee}}}}\, \\ (\,\href{Ledger.Conway.Gov.html#12875}{\htmlId{12875}{\htmlClass{Bound}{\text{new}}}}\, , \,\href{Ledger.Conway.Gov.html#12881}{\htmlId{12881}{\htmlClass{Bound}{\text{rem}}}}\, , \,\href{Ledger.Conway.Gov.html#12887}{\htmlId{12887}{\htmlClass{Bound}{\text{q}}}}\,) \end{pmatrix}$ =
  p  nothing × (∀[ e  range new ]  epoch < e) × (dom new  rem ≡ᵉ )
actionValid rewardCreds p ppolicy epoch _ =
  p  nothing

actionWellFormed : GovAction  Type
actionWellFormed $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#898}{\htmlId{13081}{\htmlClass{InductiveConstructor}{\text{ChangePParams}}}}\, \\ \,\href{Ledger.Conway.Gov.html#13097}{\htmlId{13097}{\htmlClass{Bound}{\text{x}}}}\, \end{pmatrix}$ = ppdWellFormed x
actionWellFormed $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#933}{\htmlId{13140}{\htmlClass{InductiveConstructor}{\text{TreasuryWdrl}}}}\,  \\ \,\href{Ledger.Conway.Gov.html#13156}{\htmlId{13156}{\htmlClass{Bound}{\text{x}}}}\, \end{pmatrix}$ =
  (∀[ a  dom x ] NetworkIdOf a  NetworkId) × (∃[ v  range x ] ¬ (v  0))
actionWellFormed _                 = 


actionValid? :  {rewardCreds p ppolicy epoch a}  actionValid rewardCreds p ppolicy epoch a 
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#758}{\htmlId{13396}{\htmlClass{InductiveConstructor}{\text{NoConfidence}}}}\,    \\ \,\htmlId{13414}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#793}{\htmlId{13446}{\htmlClass{InductiveConstructor}{\text{UpdateCommittee}}}}\, \\ \,\htmlId{13464}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#828}{\htmlId{13496}{\htmlClass{InductiveConstructor}{\text{NewConstitution}}}}\, \\ \,\htmlId{13514}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#863}{\htmlId{13546}{\htmlClass{InductiveConstructor}{\text{TriggerHF}}}}\,       \\ \,\htmlId{13564}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#898}{\htmlId{13596}{\htmlClass{InductiveConstructor}{\text{ChangePParams}}}}\,   \\ \,\htmlId{13614}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#933}{\htmlId{13646}{\htmlClass{InductiveConstructor}{\text{TreasuryWdrl}}}}\,    \\ \,\htmlId{13664}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionValid? {a = $\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#968}{\htmlId{13696}{\htmlClass{InductiveConstructor}{\text{Info}}}}\,            \\ \,\htmlId{13714}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it

actionWellFormed? :  {a}  actionWellFormed a 
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#758}{\htmlId{13797}{\htmlClass{InductiveConstructor}{\text{NoConfidence}}}}\,    \\ \,\htmlId{13815}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#793}{\htmlId{13848}{\htmlClass{InductiveConstructor}{\text{UpdateCommittee}}}}\, \\ \,\htmlId{13866}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#828}{\htmlId{13899}{\htmlClass{InductiveConstructor}{\text{NewConstitution}}}}\, \\ \,\htmlId{13917}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#863}{\htmlId{13950}{\htmlClass{InductiveConstructor}{\text{TriggerHF}}}}\,       \\ \,\htmlId{13968}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#898}{\htmlId{14001}{\htmlClass{InductiveConstructor}{\text{ChangePParams}}}}\,   \\ \,\htmlId{14019}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#933}{\htmlId{14052}{\htmlClass{InductiveConstructor}{\text{TreasuryWdrl}}}}\,    \\ \,\htmlId{14070}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it
actionWellFormed? {$\begin{pmatrix} \,\href{Ledger.Conway.GovernanceActions.html#968}{\htmlId{14103}{\htmlClass{InductiveConstructor}{\text{Info}}}}\,            \\ \,\htmlId{14121}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$} = it


open GovEnv
open PParams hiding (a)

variable
  machr : Maybe Anchor
  achr : Anchor
  ast  : GovActionState


data _⊢_⇀⦇_,GOV⦈_ where


  GOV-Vote :
     (aid , ast)  fromList s
     canVote (Γ .pparams) (action ast) (proj₁ voter)
     isRegistered Γ voter
     ¬ expired (Γ .epoch) ast
      ───────────────────────────────────────
      (Γ , k)  s ⇀⦇ inj₁ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.html#1961}{\htmlId{14502}{\htmlClass{Generalizable}{\text{aid}}}}\, \\ \,\href{Ledger.Conway.Gov.html#1981}{\htmlId{14508}{\htmlClass{Generalizable}{\text{voter}}}}\, \\ \,\href{Ledger.Conway.Gov.html#2014}{\htmlId{14516}{\htmlClass{Generalizable}{\text{v}}}}\, \\ \,\href{Ledger.Conway.Gov.html#14183}{\htmlId{14520}{\htmlClass{Generalizable}{\text{machr}}}}\, \end{pmatrix}$ ,GOV⦈ addVote s aid voter v

  GOV-Propose :
    let pp           = Γ .pparams
        e            = Γ .epoch
        enactState   = Γ .enactState
        rewardCreds  = Γ .rewardCreds
        prop         = record { returnAddr = addr ; action = a ; anchor = achr
                              ; policy = p ; deposit = d ; prevAction = prev }
    in
     actionWellFormed a
     actionValid rewardCreds p (Γ .ppolicy) e a
     d  pp .govActionDeposit
     validHFAction prop s enactState
     hasParent enactState s (a .gaType) prev
     NetworkIdOf addr  NetworkId
     CredentialOf addr  rewardCreds
      ───────────────────────────────────────
      (Γ , k)  s ⇀⦇ inj₂ prop ,GOV⦈ addAction s (pp .govActionLifetime +ᵉ e)
                                                 (Γ .txid , k) addr a prev

_⊢_⇀⦇_,GOVS⦈_ = ReflexiveTransitiveClosureᵢ {sts = _⊢_⇀⦇_,GOV⦈_}