{-# OPTIONS --safe #-} open import Data.Nat.Properties using (+-0-monoid) open import Data.Rational using (ℚ) open import Ledger.Prelude open import Ledger.Conway.Gov.Base module Ledger.Conway.Enact (gs : _) (open GovStructure gs) where open import Ledger.Conway.Gov.Actions gs record EnactEnv : Type where field gid : GovActionID treasury : Coin epoch : Epoch record EnactState : Type where field cc : HashProtected (Maybe ((Credential ⇀ Epoch) × ℚ)) constitution : HashProtected (DocHash × Maybe ScriptHash) pv : HashProtected ProtVer pparams : HashProtected PParams withdrawals : RwdAddr ⇀ Coin record HasEnactState {a} (A : Type a) : Type a where field EnactStateOf : A → EnactState open HasEnactState ⦃...⦄ public instance HasPParams-EnactState : HasPParams EnactState HasPParams-EnactState .PParamsOf = proj₁ ∘ EnactState.pparams HasccMaxTermLength-EnactState : HasccMaxTermLength EnactState HasccMaxTermLength-EnactState .ccMaxTermLengthOf = PParams.ccMaxTermLength ∘ PParamsOf unquoteDecl HasCast-EnactEnv = derive-HasCast [ (quote EnactEnv , HasCast-EnactEnv) ] open EnactState ccCreds : HashProtected (Maybe ((Credential ⇀ Epoch) × ℚ)) → ℙ Credential ccCreds (just x , _) = dom (x .proj₁) ccCreds (nothing , _) = ∅ getHash : ∀ {a} → NeedsHash a → Maybe GovActionID getHash {NoConfidence} h = just h getHash {UpdateCommittee} h = just h getHash {NewConstitution} h = just h getHash {TriggerHF} h = just h getHash {ChangePParams} h = just h getHash {TreasuryWdrl} _ = nothing getHash {Info} _ = nothing getHashES : EnactState → GovActionType → Maybe GovActionID getHashES es NoConfidence = just (es .cc .proj₂) getHashES es (UpdateCommittee) = just (es .cc .proj₂) getHashES es (NewConstitution) = just (es .constitution .proj₂) getHashES es (TriggerHF) = just (es .pv .proj₂) getHashES es (ChangePParams) = just (es .pparams .proj₂) getHashES es (TreasuryWdrl) = nothing getHashES es Info = nothing data _⊢_⇀⦇_,ENACT⦈_ : EnactEnv → EnactState → GovAction → EnactState → Type open EnactState private variable s : EnactState up : PParamsUpdate new : Credential ⇀ Epoch rem : ℙ Credential q : ℚ dh : DocHash sh : Maybe ScriptHash v : ProtVer wdrl : RwdAddr ⇀ Coin t : Coin gid : GovActionID e : Epoch instance _ = +-0-monoid open PParams using (ccMaxTermLength) open EnactState using (cc) data _⊢_⇀⦇_,ENACT⦈_ where Enact-NoConf : ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{2615}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{2621}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{2625}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#710}{\htmlId{2638}{\htmlClass{InductiveConstructor}{\text{NoConfidence}}}}\, \\ \,\htmlId{2653}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$ ,ENACT⦈ record s { cc = nothing , gid } Enact-UpdComm : let old = maybe proj₁ ∅ (s .cc .proj₁) maxTerm = ccMaxTermLengthOf s +ᵉ e in ∀[ term ∈ range new ] term ≤ maxTerm ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{2932}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{2938}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{2942}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#745}{\htmlId{2955}{\htmlClass{InductiveConstructor}{\text{UpdateCommittee}}}}\, \\ (\,\href{Ledger.Conway.Enact.html#2253}{\htmlId{2974}{\htmlClass{Generalizable}{\text{new}}}}\, , \,\href{Ledger.Conway.Enact.html#2280}{\htmlId{2980}{\htmlClass{Generalizable}{\text{rem}}}}\, , \,\href{Ledger.Conway.Enact.html#2301}{\htmlId{2986}{\htmlClass{Generalizable}{\text{q}}}}\,) \end{pmatrix}$ ,ENACT⦈ record s { cc = just ((new ∪ˡ old) ∣ rem ᶜ , q) , gid } Enact-NewConst : ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{3133}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{3139}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{3143}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#780}{\htmlId{3156}{\htmlClass{InductiveConstructor}{\text{NewConstitution}}}}\, \\ (\,\href{Ledger.Conway.Enact.html#2309}{\htmlId{3175}{\htmlClass{Generalizable}{\text{dh}}}}\, , \,\href{Ledger.Conway.Enact.html#2324}{\htmlId{3180}{\htmlClass{Generalizable}{\text{sh}}}}\,) \end{pmatrix}$ ,ENACT⦈ record s { constitution = (dh , sh) , gid } Enact-HF : ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{3305}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{3311}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{3315}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#815}{\htmlId{3328}{\htmlClass{InductiveConstructor}{\text{TriggerHF}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2348}{\htmlId{3340}{\htmlClass{Generalizable}{\text{v}}}}\, \end{pmatrix}$ ,ENACT⦈ record s { pv = v , gid } Enact-PParams : ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{3449}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{3455}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{3459}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#850}{\htmlId{3472}{\htmlClass{InductiveConstructor}{\text{ChangePParams}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2232}{\htmlId{3488}{\htmlClass{Generalizable}{\text{up}}}}\, \end{pmatrix}$ ,ENACT⦈ record s { pparams = applyUpdate (PParamsOf s) up , gid } Enact-Wdrl : let newWdrls = s .withdrawals ∪⁺ wdrl in ∑[ x ← newWdrls ] x ≤ t ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{3702}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{3708}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{3712}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#885}{\htmlId{3725}{\htmlClass{InductiveConstructor}{\text{TreasuryWdrl}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2362}{\htmlId{3740}{\htmlClass{Generalizable}{\text{wdrl}}}}\, \end{pmatrix}$ ,ENACT⦈ record s { withdrawals = newWdrls } Enact-Info : ─────────────────────────────────────── $\begin{pmatrix} \,\href{Ledger.Conway.Enact.html#2397}{\htmlId{3859}{\htmlClass{Generalizable}{\text{gid}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2386}{\htmlId{3865}{\htmlClass{Generalizable}{\text{t}}}}\, \\ \,\href{Ledger.Conway.Enact.html#2417}{\htmlId{3869}{\htmlClass{Generalizable}{\text{e}}}}\, \end{pmatrix}$ ⊢ s ⇀⦇ $\begin{pmatrix} \,\href{Ledger.Conway.Gov.Actions.html#920}{\htmlId{3882}{\htmlClass{InductiveConstructor}{\text{Info}}}}\, \\ \,\htmlId{3889}{\htmlClass{Symbol}{\text{\_}}}\, \end{pmatrix}$ ,ENACT⦈ s