diff options
author | Greg Brown <greg.brown@cl.cam.ac.uk> | 2022-04-18 15:05:24 +0100 |
---|---|---|
committer | Greg Brown <greg.brown@cl.cam.ac.uk> | 2022-04-18 15:21:38 +0100 |
commit | 00a0ce9082b4cc1389815defcc806efd4a9b80f4 (patch) | |
tree | aebdc99b954be177571697fc743ee75841c98b2e | |
parent | 24845ef25e12864711552ebc75a1f54903bee50c (diff) |
Do a big refactor.
- Replace the decidable predicates on expressions and statements with
separate data types.
- Reorganise the Hoare logic semantics to remove unnecessary
definitions.
- Make liberal use of modules to group related definitions together.
- Unify the types for denotational and Hoare logic semantics.
- Make bits an abstraction of array types.
-rw-r--r-- | Everything.agda | 9 | ||||
-rw-r--r-- | src/Helium/Data/Pseudocode/Core.agda | 522 | ||||
-rw-r--r-- | src/Helium/Data/Pseudocode/Manipulate.agda | 2756 | ||||
-rw-r--r-- | src/Helium/Data/Pseudocode/Properties.agda | 109 | ||||
-rw-r--r-- | src/Helium/Instructions/Base.agda | 240 | ||||
-rw-r--r-- | src/Helium/Instructions/Instances/Barrett.agda | 6 | ||||
-rw-r--r-- | src/Helium/Semantics/Axiomatic.agda | 47 | ||||
-rw-r--r-- | src/Helium/Semantics/Axiomatic/Assertion.agda | 247 | ||||
-rw-r--r-- | src/Helium/Semantics/Axiomatic/Core.agda | 85 | ||||
-rw-r--r-- | src/Helium/Semantics/Axiomatic/Term.agda | 873 | ||||
-rw-r--r-- | src/Helium/Semantics/Axiomatic/Triple.agda | 61 | ||||
-rw-r--r-- | src/Helium/Semantics/Core.agda | 209 | ||||
-rw-r--r-- | src/Helium/Semantics/Denotational/Core.agda | 416 |
13 files changed, 2658 insertions, 2922 deletions
diff --git a/Everything.agda b/Everything.agda index 2f76af0..87109e2 100644 --- a/Everything.agda +++ b/Everything.agda @@ -96,9 +96,6 @@ import Helium.Data.Pseudocode.Core -- Ways to modify pseudocode statements and expressions. import Helium.Data.Pseudocode.Manipulate --- Basic properties of the pseudocode data types -import Helium.Data.Pseudocode.Properties - -- Definition of instructions using the Armv8-M pseudocode. import Helium.Instructions.Base @@ -117,15 +114,15 @@ import Helium.Semantics.Axiomatic -- Definition of assertions used in correctness triples import Helium.Semantics.Axiomatic.Assertion --- Base definitions for the axiomatic semantics -import Helium.Semantics.Axiomatic.Core - -- Definition of terms for use in assertions import Helium.Semantics.Axiomatic.Term -- Definition of Hoare triples import Helium.Semantics.Axiomatic.Triple +-- Base definitions for semantics +import Helium.Semantics.Core + -- Base definitions for the denotational semantics. import Helium.Semantics.Denotational.Core diff --git a/src/Helium/Data/Pseudocode/Core.agda b/src/Helium/Data/Pseudocode/Core.agda index 079e2ce..579d68d 100644 --- a/src/Helium/Data/Pseudocode/Core.agda +++ b/src/Helium/Data/Pseudocode/Core.agda @@ -11,20 +11,18 @@ module Helium.Data.Pseudocode.Core where open import Data.Bool using (Bool; true; false) open import Data.Fin as Fin using (Fin) open import Data.Fin.Patterns +open import Data.Integer as ℤ using (ℤ) open import Data.Nat as ℕ using (ℕ; zero; suc) open import Data.Nat.Properties using (+-comm) -open import Data.Product using (∃; _,_; proj₂; uncurry) -open import Data.Sum using ([_,_]′; inj₁; inj₂) +open import Data.Product using (_×_; _,_) +open import Data.Unit using (⊤) open import Data.Vec using (Vec; []; _∷_; lookup; map) -open import Data.Vec.Relation.Unary.All using (All; []; _∷_; reduce) -open import Data.Vec.Relation.Unary.Any using (Any; here; there) -open import Function as F using (_∘_; _∘′_; _∋_) -open import Relation.Binary.PropositionalEquality using (_≡_; refl) -open import Relation.Nullary using (Dec; yes; no) -open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromWitness; map′) -open import Relation.Nullary.Product using (_×-dec_) -open import Relation.Nullary.Sum using (_⊎-dec_) -open import Relation.Unary using (Decidable) +open import Data.Vec.Relation.Unary.All using (All; []; _∷_) +open import Relation.Binary.PropositionalEquality using (_≡_) + +private + variable + k m n o : ℕ --- The set of types and boolean properties of them data Type : Set where @@ -33,297 +31,221 @@ data Type : Set where fin : (n : ℕ) → Type real : Type bit : Type - bits : (n : ℕ) → Type - tuple : ∀ n → Vec Type n → Type + tuple : Vec Type n → Type array : Type → (n : ℕ) → Type +bits : ℕ → Type +bits = array bit + +private + variable + t t′ t₁ t₂ : Type + Σ Γ Δ ts : Vec Type n + data HasEquality : Type → Set where - bool : HasEquality bool - int : HasEquality int - fin : ∀ {n} → HasEquality (fin n) - real : HasEquality real - bit : HasEquality bit - bits : ∀ {n} → HasEquality (bits n) - -hasEquality? : Decidable HasEquality -hasEquality? bool = yes bool -hasEquality? int = yes int -hasEquality? (fin n) = yes fin -hasEquality? real = yes real -hasEquality? bit = yes bit -hasEquality? (bits n) = yes bits -hasEquality? (tuple n x) = no (λ ()) -hasEquality? (array t n) = no (λ ()) + instance bool : HasEquality bool + instance int : HasEquality int + instance fin : HasEquality (fin n) + instance real : HasEquality real + instance bit : HasEquality bit + instance array : ∀ {t n} → ⦃ HasEquality t ⦄ → HasEquality (array t n) + +data Ordered : Type → Set where + instance int : Ordered int + instance fin : Ordered (fin n) + instance real : Ordered real data IsNumeric : Type → Set where - int : IsNumeric int - real : IsNumeric real - -isNumeric? : Decidable IsNumeric -isNumeric? bool = no (λ ()) -isNumeric? int = yes int -isNumeric? (fin n) = no (λ ()) -isNumeric? real = yes real -isNumeric? bit = no (λ ()) -isNumeric? (bits n) = no (λ ()) -isNumeric? (tuple n x) = no (λ ()) -isNumeric? (array t n) = no (λ ()) - -combineNumeric : ∀ t₁ t₂ → (isNumeric₁ : True (isNumeric? t₁)) → (isNumeric₂ : True (isNumeric? t₂)) → Type -combineNumeric int int _ _ = int -combineNumeric int real _ _ = real -combineNumeric real _ _ _ = real - -data Sliced : Set where - bits : Sliced - array : Type → Sliced - -asType : Sliced → ℕ → Type -asType bits n = bits n -asType (array t) n = array t n - -elemType : Sliced → Type -elemType bits = bit -elemType (array t) = t - ---- Literals - -data Literal : Type → Set where - _′b : Bool → Literal bool - _′i : ℕ → Literal int - _′f : ∀ {n} → Fin n → Literal (fin n) - _′r : ℕ → Literal real - _′x : Bool → Literal bit - _′xs : ∀ {n} → Vec Bool n → Literal (bits n) - _′a : ∀ {n t} → Literal t → Literal (array t n) - ---- Expressions, references, statements, functions and procedures - -module Code {o} (Σ : Vec Type o) where - data Expression {n} (Γ : Vec Type n) : Type → Set - data CanAssign {n Γ} : ∀ {t} → Expression {n} Γ t → Set - canAssign? : ∀ {n Γ t} → Decidable (CanAssign {n} {Γ} {t}) - data ReferencesState {n Γ} : ∀ {t} → Expression {n} Γ t → Set - referencesState? : ∀ {n Γ t} → Decidable (ReferencesState {n} {Γ} {t}) - data Statement {n} (Γ : Vec Type n) : Set - data ChangesState {n Γ} : Statement {n} Γ → Set - changesState? : ∀ {n Γ} → Decidable (ChangesState {n} {Γ}) - data Function {n} (Γ : Vec Type n) (ret : Type) : Set - data Procedure {n} (Γ : Vec Type n) : Set - - infix 8 -_ - infixr 7 _^_ - infixl 6 _*_ _and_ _>>_ - -- infixl 6 _/_ - infixl 5 _+_ _or_ _&&_ _||_ - infix 4 _≟_ _<?_ - - data Expression {n} Γ where - lit : ∀ {t} → Literal t → Expression Γ t - state : ∀ i → Expression Γ (lookup Σ i) - var : ∀ i → Expression Γ (lookup Γ i) - abort : ∀ {t} → Expression Γ (fin 0) → Expression Γ t - _≟_ : ∀ {t} {hasEquality : True (hasEquality? t)} → Expression Γ t → Expression Γ t → Expression Γ bool - _<?_ : ∀ {t} {isNumeric : True (isNumeric? t)} → Expression Γ t → Expression Γ t → Expression Γ bool - inv : Expression Γ bool → Expression Γ bool - _&&_ : Expression Γ bool → Expression Γ bool → Expression Γ bool - _||_ : Expression Γ bool → Expression Γ bool → Expression Γ bool - not : ∀ {m} → Expression Γ (bits m) → Expression Γ (bits m) - _and_ : ∀ {m} → Expression Γ (bits m) → Expression Γ (bits m) → Expression Γ (bits m) - _or_ : ∀ {m} → Expression Γ (bits m) → Expression Γ (bits m) → Expression Γ (bits m) - [_] : ∀ {t} → Expression Γ (elemType t) → Expression Γ (asType t 1) - unbox : ∀ {t} → Expression Γ (asType t 1) → Expression Γ (elemType t) - splice : ∀ {m n t} → Expression Γ (asType t m) → Expression Γ (asType t n) → Expression Γ (fin (suc n)) → Expression Γ (asType t (n ℕ.+ m)) - cut : ∀ {m n t} → Expression Γ (asType t (n ℕ.+ m)) → Expression Γ (fin (suc n)) → Expression Γ (tuple 2 (asType t m ∷ asType t n ∷ [])) - cast : ∀ {i j t} → .(eq : i ≡ j) → Expression Γ (asType t i) → Expression Γ (asType t j) - -_ : ∀ {t} {isNumeric : True (isNumeric? t)} → Expression Γ t → Expression Γ t - _+_ : ∀ {t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ isNumeric₁ isNumeric₂) - _*_ : ∀ {t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ isNumeric₁ isNumeric₂) - -- _/_ : Expression Γ real → Expression Γ real → Expression Γ real - _^_ : ∀ {t} {isNumeric : True (isNumeric? t)} → Expression Γ t → ℕ → Expression Γ t - _>>_ : Expression Γ int → ℕ → Expression Γ int - rnd : Expression Γ real → Expression Γ int - fin : ∀ {k ms n} → (All (Fin) ms → Fin n) → Expression Γ (tuple k (map fin ms)) → Expression Γ (fin n) - asInt : ∀ {m} → Expression Γ (fin m) → Expression Γ int - nil : Expression Γ (tuple 0 []) - cons : ∀ {m t ts} → Expression Γ t → Expression Γ (tuple m ts) → Expression Γ (tuple (suc m) (t ∷ ts)) - head : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ t - tail : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ (tuple m ts) - call : ∀ {t m Δ} → Function Δ t → All (Expression Γ) {m} Δ → Expression Γ t - if_then_else_ : ∀ {t} → Expression Γ bool → Expression Γ t → Expression Γ t → Expression Γ t - - data CanAssign {n} {Γ} where - state : ∀ i → CanAssign (state i) - var : ∀ i → CanAssign (var i) - abort : ∀ {t} e → CanAssign (abort {t = t} e) - [_] : ∀ {t e} → CanAssign e → CanAssign ([_] {t = t} e) - unbox : ∀ {t e} → CanAssign e → CanAssign (unbox {t = t} e) - splice : ∀ {m n t e₁ e₂} → CanAssign e₁ → CanAssign e₂ → ∀ e₃ → CanAssign (splice {m = m} {n} {t} e₁ e₂ e₃) - cut : ∀ {m n t e₁} → CanAssign e₁ → ∀ e₂ → CanAssign (cut {m = m} {n} {t} e₁ e₂) - cast : ∀ {i j t e} .(eq : i ≡ j) → CanAssign e → CanAssign (cast {t = t} eq e) - nil : CanAssign nil - cons : ∀ {m t ts e₁ e₂} → CanAssign e₁ → CanAssign e₂ → CanAssign (cons {m = m} {t} {ts} e₁ e₂) - head : ∀ {m t ts e} → CanAssign e → CanAssign (head {m = m} {t} {ts} e) - tail : ∀ {m t ts e} → CanAssign e → CanAssign (tail {m = m} {t} {ts} e) - - canAssign? (lit x) = no λ () - canAssign? (state i) = yes (state i) - canAssign? (var i) = yes (var i) - canAssign? (abort e) = yes (abort e) - canAssign? (e ≟ e₁) = no λ () - canAssign? (e <? e₁) = no λ () - canAssign? (inv e) = no λ () - canAssign? (e && e₁) = no λ () - canAssign? (e || e₁) = no λ () - canAssign? (not e) = no λ () - canAssign? (e and e₁) = no λ () - canAssign? (e or e₁) = no λ () - canAssign? [ e ] = map′ [_] (λ { [ e ] → e }) (canAssign? e) - canAssign? (unbox e) = map′ unbox (λ { (unbox e) → e }) (canAssign? e) - canAssign? (splice e e₁ e₂) = map′ (λ (e , e₁) → splice e e₁ e₂) (λ { (splice e e₁ _) → e , e₁ }) (canAssign? e ×-dec canAssign? e₁) - canAssign? (cut e e₁) = map′ (λ e → cut e e₁) (λ { (cut e e₁) → e }) (canAssign? e) - canAssign? (cast eq e) = map′ (cast eq) (λ { (cast eq e) → e }) (canAssign? e) - canAssign? (- e) = no λ () - canAssign? (e + e₁) = no λ () - canAssign? (e * e₁) = no λ () - -- canAssign? (e / e₁) = no λ () - canAssign? (e ^ e₁) = no λ () - canAssign? (e >> e₁) = no λ () - canAssign? (rnd e) = no λ () - canAssign? (fin x e) = no λ () - canAssign? (asInt e) = no λ () - canAssign? nil = yes nil - canAssign? (cons e e₁) = map′ (uncurry cons) (λ { (cons e e₁) → e , e₁ }) (canAssign? e ×-dec canAssign? e₁) - canAssign? (head e) = map′ head (λ { (head e) → e }) (canAssign? e) - canAssign? (tail e) = map′ tail (λ { (tail e) → e }) (canAssign? e) - canAssign? (call x e) = no λ () - canAssign? (if e then e₁ else e₂) = no λ () - - data ReferencesState where - state : ∀ i → ReferencesState (state i) - [_] : ∀ {t e} → ReferencesState e → ReferencesState ([_] {t = t} e) - unbox : ∀ {t e} → ReferencesState e → ReferencesState (unbox {t = t} e) - spliceˡ : ∀ {m n t e} → ReferencesState e → ∀ e₁ e₂ → ReferencesState (splice {m = m} {n} {t} e e₁ e₂) - spliceʳ : ∀ {m n t} e {e₁} → ReferencesState e₁ → ∀ e₂ → ReferencesState (splice {m = m} {n} {t} e e₁ e₂) - cut : ∀ {m n t e} → ReferencesState e → ∀ e₁ → ReferencesState (cut {m = m} {n} {t} e e₁) - cast : ∀ {i j t} .(eq : i ≡ j) {e} → ReferencesState e → ReferencesState (cast {t = t} eq e) - consˡ : ∀ {m t ts e} → ReferencesState e → ∀ e₁ → ReferencesState (cons {m = m} {t} {ts} e e₁) - consʳ : ∀ {m t ts} e {e₁} → ReferencesState e₁ → ReferencesState (cons {m = m} {t} {ts} e e₁) - head : ∀ {m t ts e} → ReferencesState e → ReferencesState (head {m = m} {t} {ts} e) - tail : ∀ {m t ts e} → ReferencesState e → ReferencesState (tail {m = m} {t} {ts} e) - - referencesState? (lit x) = no λ () - referencesState? (state i) = yes (state i) - referencesState? (var i) = no λ () - referencesState? (abort e) = no λ () - referencesState? (e ≟ e₁) = no λ () - referencesState? (e <? e₁) = no λ () - referencesState? (inv e) = no λ () - referencesState? (e && e₁) = no λ () - referencesState? (e || e₁) = no λ () - referencesState? (not e) = no λ () - referencesState? (e and e₁) = no λ () - referencesState? (e or e₁) = no λ () - referencesState? [ e ] = map′ [_] (λ { [ e ] → e }) (referencesState? e) - referencesState? (unbox e) = map′ unbox (λ { (unbox e) → e }) (referencesState? e) - referencesState? (splice e e₁ e₂) = map′ [ (λ e → spliceˡ e e₁ e₂) , (λ e₁ → spliceʳ e e₁ e₂) ]′ (λ { (spliceˡ e e₁ e₂) → inj₁ e ; (spliceʳ e e₁ e₂) → inj₂ e₁ }) (referencesState? e ⊎-dec referencesState? e₁) - referencesState? (cut e e₁) = map′ (λ e → cut e e₁) (λ { (cut e e₁) → e }) (referencesState? e) - referencesState? (cast eq e) = map′ (cast eq) (λ { (cast eq e) → e }) (referencesState? e) - referencesState? (- e) = no λ () - referencesState? (e + e₁) = no λ () - referencesState? (e * e₁) = no λ () - referencesState? (e ^ x) = no λ () - referencesState? (e >> x) = no λ () - referencesState? (rnd e) = no λ () - referencesState? (fin x e) = no λ () - referencesState? (asInt e) = no λ () - referencesState? nil = no λ () - referencesState? (cons e e₁) = map′ [ (λ e → consˡ e e₁) , (λ e₁ → consʳ e e₁) ]′ (λ { (consˡ e e₁) → inj₁ e ; (consʳ e e₁) → inj₂ e₁ }) (referencesState? e ⊎-dec referencesState? e₁) - referencesState? (head e) = map′ head (λ { (head e) → e }) (referencesState? e) - referencesState? (tail e) = map′ tail (λ { (tail e) → e }) (referencesState? e) - referencesState? (call f es) = no λ () - referencesState? (if e then e₁ else e₂) = no λ () - - infix 4 _≔_ - infixl 2 if_then_else_ if_then_ - infixl 1 _∙_ _∙return_ - infix 1 _∙end - - data Statement Γ where - _∙_ : Statement Γ → Statement Γ → Statement Γ - skip : Statement Γ - _≔_ : ∀ {t} → (ref : Expression Γ t) → {canAssign : True (canAssign? ref)} → Expression Γ t → Statement Γ - declare : ∀ {t} → Expression Γ t → Statement (t ∷ Γ) → Statement Γ - invoke : ∀ {m Δ} → Procedure Δ → All (Expression Γ) {m} Δ → Statement Γ - if_then_ : Expression Γ bool → Statement Γ → Statement Γ - if_then_else_ : Expression Γ bool → Statement Γ → Statement Γ → Statement Γ - for : ∀ m → Statement (fin m ∷ Γ) → Statement Γ - - data ChangesState where - _∙ˡ_ : ∀ {s} → ChangesState s → ∀ s₁ → ChangesState (s ∙ s₁) - _∙ʳ_ : ∀ s {s₁} → ChangesState s₁ → ChangesState (s ∙ s₁) - _≔_ : ∀ {t} ref {canAssign : True (canAssign? ref)} {refsState : True (referencesState? ref)} e₂ → ChangesState (_≔_ {t = t} ref {canAssign} e₂) - declare : ∀ {t} e {s} → ChangesState s → ChangesState (declare {t = t} e s) - invoke : ∀ {m Δ} p es → ChangesState (invoke {m = m} {Δ} p es) - if_then_ : ∀ e {s} → ChangesState s → ChangesState (if e then s) - if_then′_else_ : ∀ e {s} → ChangesState s → ∀ s₁ → ChangesState (if e then s else s₁) - if_then_else′_ : ∀ e s {s₁} → ChangesState s₁ → ChangesState (if e then s else s₁) - for : ∀ m {s} → ChangesState s → ChangesState (for m s) - - changesState? (s ∙ s₁) = map′ [ _∙ˡ s₁ , s ∙ʳ_ ]′ (λ { (s ∙ˡ s₁) → inj₁ s ; (s ∙ʳ s₁) → inj₂ s₁ }) (changesState? s ⊎-dec changesState? s₁) - changesState? skip = no λ () - changesState? (_≔_ ref e) = map′ (λ refsState → _≔_ ref {refsState = fromWitness refsState} e) (λ { (_≔_ ref {refsState = refsState} e) → toWitness refsState }) (referencesState? ref) - changesState? (declare e s) = map′ (declare e) (λ { (declare e s) → s }) (changesState? s) - changesState? (invoke p e) = yes (invoke p e) - changesState? (if e then s) = map′ (if e then_) (λ { (if e then s) → s }) (changesState? s) - changesState? (if e then s else s₁) = map′ [ if e then′_else s₁ , if e then s else′_ ]′ (λ { (if e then′ s else s₁) → inj₁ s ; (if e then s else′ s₁) → inj₂ s₁ }) (changesState? s ⊎-dec changesState? s₁) - changesState? (for m s) = map′ (for m) (λ { (for m s) → s }) (changesState? s) - - data Function Γ ret where - _∙return_ : (s : Statement Γ) → {False (changesState? s)} → Expression Γ ret → Function Γ ret - declare : ∀ {t} → Expression Γ t → Function (t ∷ Γ) ret → Function Γ ret - - data Procedure Γ where - _∙end : Statement Γ → Procedure Γ - - infixl 6 _<<_ - infixl 5 _-_ _∶_ - - tup : ∀ {n Γ m ts} → All (Expression {n} Γ) ts → Expression Γ (tuple m ts) - tup [] = nil - tup (e ∷ es) = cons e (tup es) - - _∶_ : ∀ {n Γ i j t} → Expression {n} Γ (asType t j) → Expression Γ (asType t i) → Expression Γ (asType t (i ℕ.+ j)) - e₁ ∶ e₂ = splice e₁ e₂ (lit (Fin.fromℕ _ ′f)) - - slice : ∀ {n Γ i j t} → Expression {n} Γ (asType t (i ℕ.+ j)) → Expression Γ (fin (suc i)) → Expression Γ (asType t j) - slice e₁ e₂ = head (cut e₁ e₂) - - slice′ : ∀ {n Γ i j t} → Expression {n} Γ (asType t (i ℕ.+ j)) → Expression Γ (fin (suc j)) → Expression Γ (asType t i) - slice′ {i = i} e₁ e₂ = slice (cast (+-comm i _) e₁) e₂ - - _-_ : ∀ {n Γ t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression {n} Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ isNumeric₁ isNumeric₂) - _-_ {isNumeric₂ = isNumeric₂} x y = x + (-_ {isNumeric = isNumeric₂} y) - - _<<_ : ∀ {n Γ} → Expression {n} Γ int → ℕ → Expression Γ int - x << n = rnd (x * lit (2 ′r) ^ n) - - get : ∀ {n Γ} → ℕ → Expression {n} Γ int → Expression Γ bit - get i x = if x - x >> suc i << suc i <? lit (2 ′i) ^ i then lit (false ′x) else lit (true ′x) - - uint : ∀ {n Γ m} → Expression {n} Γ (bits m) → Expression Γ int - uint {m = zero} x = lit (0 ′i) - uint {m = suc m} x = - lit (2 ′i) * uint {m = m} (slice x (lit (1F ′f))) + - ( if slice′ x (lit (0F ′f)) ≟ lit ((true ∷ []) ′xs) - then lit (1 ′i) - else lit (0 ′i)) - - sint : ∀ {n Γ m} → Expression {n} Γ (bits m) → Expression Γ int - sint {m = zero} x = lit (0 ′i) - sint {m = suc zero} x = if x ≟ lit ((true ∷ []) ′xs) then - lit (1 ′i) else lit (0 ′i) - sint {m = suc (suc m)} x = - lit (2 ′i) * sint (slice {i = 1} x (lit (1F ′f))) + - ( if slice′ x (lit (0F ′f)) ≟ lit ((true ∷ []) ′xs) - then lit (1 ′i) - else lit (0 ′i)) + instance int : IsNumeric int + instance real : IsNumeric real + +_+ᵗ_ : IsNumeric t₁ → IsNumeric t₂ → Type +int +ᵗ int = int +int +ᵗ real = real +real +ᵗ t₂ = real + +literalType : Type → Set +literalTypes : Vec Type n → Set + +literalType bool = Bool +literalType int = ℤ +literalType (fin n) = Fin n +literalType real = ℤ +literalType bit = Bool +literalType (tuple ts) = literalTypes ts +literalType (array t n) = Vec (literalType t) n + +literalTypes [] = ⊤ +literalTypes (t ∷ []) = literalType t +literalTypes (t ∷ t′ ∷ ts) = literalType t × literalTypes (t′ ∷ ts) + +infix 8 -_ +infixr 7 _^_ +infixl 6 _*_ _and_ _>>_ +infixl 5 _+_ _or_ _&&_ _||_ +infix 4 _≟_ _<?_ +infix 3 _≔_ + +infixl 2 if_then_ if_then_else_ +infixl 1 _∙_ _∙return_ +infix 1 _∙end + +data Expression (Σ : Vec Type o) (Γ : Vec Type n) : Type → Set +data Reference (Σ : Vec Type o) (Γ : Vec Type n) : Type → Set +data LocalReference (Σ : Vec Type o) (Γ : Vec Type n) : Type → Set +data Statement (Σ : Vec Type o) (Γ : Vec Type n) : Set +data LocalStatement (Σ : Vec Type o) (Γ : Vec Type n) : Set +data Function (Σ : Vec Type o) (Γ : Vec Type n) (ret : Type) : Set +data Procedure (Σ : Vec Type o) (Γ : Vec Type n) : Set + +data Expression Σ Γ where + lit : literalType t → Expression Σ Γ t + state : ∀ i → Expression Σ Γ (lookup Σ i) + var : ∀ i → Expression Σ Γ (lookup Γ i) + _≟_ : ⦃ HasEquality t ⦄ → Expression Σ Γ t → Expression Σ Γ t → Expression Σ Γ bool + _<?_ : ⦃ Ordered t ⦄ → Expression Σ Γ t → Expression Σ Γ t → Expression Σ Γ bool + inv : Expression Σ Γ bool → Expression Σ Γ bool + _&&_ : Expression Σ Γ bool → Expression Σ Γ bool → Expression Σ Γ bool + _||_ : Expression Σ Γ bool → Expression Σ Γ bool → Expression Σ Γ bool + not : Expression Σ Γ (bits n) → Expression Σ Γ (bits n) + _and_ : Expression Σ Γ (bits n) → Expression Σ Γ (bits n) → Expression Σ Γ (bits n) + _or_ : Expression Σ Γ (bits n) → Expression Σ Γ (bits n) → Expression Σ Γ (bits n) + [_] : Expression Σ Γ t → Expression Σ Γ (array t 1) + unbox : Expression Σ Γ (array t 1) → Expression Σ Γ t + merge : Expression Σ Γ (array t m) → Expression Σ Γ (array t n) → Expression Σ Γ (fin (suc n)) → Expression Σ Γ (array t (n ℕ.+ m)) + slice : Expression Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → Expression Σ Γ (array t m) + cut : Expression Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → Expression Σ Γ (array t n) + cast : .(eq : m ≡ n) → Expression Σ Γ (array t m) → Expression Σ Γ (array t n) + -_ : ⦃ IsNumeric t ⦄ → Expression Σ Γ t → Expression Σ Γ t + _+_ : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → Expression Σ Γ t₁ → Expression Σ Γ t₂ → Expression Σ Γ (isNum₁ +ᵗ isNum₂) + _*_ : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → Expression Σ Γ t₁ → Expression Σ Γ t₂ → Expression Σ Γ (isNum₁ +ᵗ isNum₂) + _^_ : ⦃ IsNumeric t ⦄ → Expression Σ Γ t → ℕ → Expression Σ Γ t + _>>_ : Expression Σ Γ int → (n : ℕ) → Expression Σ Γ int + rnd : Expression Σ Γ real → Expression Σ Γ int + fin : ∀ {ms} (f : literalTypes (map fin ms) → Fin n) → Expression Σ Γ (tuple {n = k} (map fin ms)) → Expression Σ Γ (fin n) + asInt : Expression Σ Γ (fin n) → Expression Σ Γ int + nil : Expression Σ Γ (tuple []) + cons : Expression Σ Γ t → Expression Σ Γ (tuple ts) → Expression Σ Γ (tuple (t ∷ ts)) + head : Expression Σ Γ (tuple (t ∷ ts)) → Expression Σ Γ t + tail : Expression Σ Γ (tuple (t ∷ ts)) → Expression Σ Γ (tuple ts) + call : (f : Function Σ Δ t) → All (Expression Σ Γ) Δ → Expression Σ Γ t + if_then_else_ : Expression Σ Γ bool → Expression Σ Γ t → Expression Σ Γ t → Expression Σ Γ t + +data Reference Σ Γ where + state : ∀ i → Reference Σ Γ (lookup Σ i) + var : ∀ i → Reference Σ Γ (lookup Γ i) + [_] : Reference Σ Γ t → Reference Σ Γ (array t 1) + unbox : Reference Σ Γ (array t 1) → Reference Σ Γ t + merge : Reference Σ Γ (array t m) → Reference Σ Γ (array t n) → Expression Σ Γ (fin (suc n)) → Reference Σ Γ (array t (n ℕ.+ m)) + slice : Reference Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → Reference Σ Γ (array t m) + cut : Reference Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → Reference Σ Γ (array t n) + cast : .(eq : m ≡ n) → Reference Σ Γ (array t m) → Reference Σ Γ (array t n) + nil : Reference Σ Γ (tuple []) + cons : Reference Σ Γ t → Reference Σ Γ (tuple ts) → Reference Σ Γ (tuple (t ∷ ts)) + head : Reference Σ Γ (tuple (t ∷ ts)) → Reference Σ Γ t + tail : Reference Σ Γ (tuple (t ∷ ts)) → Reference Σ Γ (tuple ts) + +data LocalReference Σ Γ where + var : ∀ i → LocalReference Σ Γ (lookup Γ i) + [_] : LocalReference Σ Γ t → LocalReference Σ Γ (array t 1) + unbox : LocalReference Σ Γ (array t 1) → LocalReference Σ Γ t + merge : LocalReference Σ Γ (array t m) → LocalReference Σ Γ (array t n) → Expression Σ Γ (fin (suc n)) → LocalReference Σ Γ (array t (n ℕ.+ m)) + slice : LocalReference Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → LocalReference Σ Γ (array t m) + cut : LocalReference Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc n)) → LocalReference Σ Γ (array t n) + cast : .(eq : m ≡ n) → LocalReference Σ Γ (array t m) → LocalReference Σ Γ (array t n) + nil : LocalReference Σ Γ (tuple []) + cons : LocalReference Σ Γ t → LocalReference Σ Γ (tuple ts) → LocalReference Σ Γ (tuple (t ∷ ts)) + head : LocalReference Σ Γ (tuple (t ∷ ts)) → LocalReference Σ Γ t + tail : LocalReference Σ Γ (tuple (t ∷ ts)) → LocalReference Σ Γ (tuple ts) + +data Statement Σ Γ where + _∙_ : Statement Σ Γ → Statement Σ Γ → Statement Σ Γ + skip : Statement Σ Γ + _≔_ : Reference Σ Γ t → Expression Σ Γ t → Statement Σ Γ + declare : Expression Σ Γ t → Statement Σ (t ∷ Γ) → Statement Σ Γ + invoke : (f : Procedure Σ Δ) → All (Expression Σ Γ) Δ → Statement Σ Γ + if_then_ : Expression Σ Γ bool → Statement Σ Γ → Statement Σ Γ + if_then_else_ : Expression Σ Γ bool → Statement Σ Γ → Statement Σ Γ → Statement Σ Γ + for : ∀ n → Statement Σ (fin n ∷ Γ) → Statement Σ Γ + +data LocalStatement Σ Γ where + _∙_ : LocalStatement Σ Γ → LocalStatement Σ Γ → LocalStatement Σ Γ + skip : LocalStatement Σ Γ + _≔_ : LocalReference Σ Γ t → Expression Σ Γ t → LocalStatement Σ Γ + declare : Expression Σ Γ t → LocalStatement Σ (t ∷ Γ) → LocalStatement Σ Γ + if_then_ : Expression Σ Γ bool → LocalStatement Σ Γ → LocalStatement Σ Γ + if_then_else_ : Expression Σ Γ bool → LocalStatement Σ Γ → LocalStatement Σ Γ → LocalStatement Σ Γ + for : ∀ n → LocalStatement Σ (fin n ∷ Γ) → LocalStatement Σ Γ + +data Function Σ Γ ret where + declare : Expression Σ Γ t → Function Σ (t ∷ Γ) ret → Function Σ Γ ret + _∙return_ : LocalStatement Σ Γ → Expression Σ Γ ret → Function Σ Γ ret + +data Procedure Σ Γ where + _∙end : Statement Σ Γ → Procedure Σ Γ + +infixl 6 _<<_ +infixl 5 _-_ _∶_ + +tup : All (Expression Σ Γ) ts → Expression Σ Γ (tuple ts) +tup [] = nil +tup (e ∷ es) = cons e (tup es) + +_∶_ : Expression Σ Γ (array t m) → Expression Σ Γ (array t n) → Expression Σ Γ (array t (n ℕ.+ m)) +e ∶ e₁ = merge e e₁ (lit (Fin.fromℕ _)) + +slice′ : Expression Σ Γ (array t (n ℕ.+ m)) → Expression Σ Γ (fin (suc m)) → Expression Σ Γ (array t n) +slice′ {m = m} e = slice (cast (+-comm _ m) e) + +_-_ : Expression Σ Γ t₁ → ⦃ isNum₁ : IsNumeric t₁ ⦄ → Expression Σ Γ t₂ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → Expression Σ Γ (isNum₁ +ᵗ isNum₂) +e - e₁ = e + (- e₁) + +_<<_ : Expression Σ Γ int → (n : ℕ) → Expression Σ Γ int +e << n = e * lit (ℤ.+ (2 ℕ.^ n)) + +getBit : ℕ → Expression Σ Γ int → Expression Σ Γ bit +getBit i x = if x - x >> suc i << suc i <? lit (ℤ.+ (2 ℕ.^ i)) then lit false else lit true + +uint : Expression Σ Γ (bits m) → Expression Σ Γ int +uint {m = 0} x = lit ℤ.0ℤ +uint {m = suc m} x = + lit (ℤ.+ 2) * uint {m = m} (slice x (lit 1F)) + + ( if slice′ x (lit 0F) ≟ lit (true ∷ []) + then lit ℤ.1ℤ + else lit ℤ.0ℤ) + +sint : Expression Σ Γ (bits m) → Expression Σ Γ int +sint {m = 0} x = lit ℤ.0ℤ +sint {m = 1} x = if x ≟ lit (true ∷ []) then lit ℤ.-1ℤ else lit ℤ.0ℤ +sint {m = suc (suc m)} x = + lit (ℤ.+ 2) * sint {m = m} (slice x (lit 1F)) + + ( if slice′ x (lit 0F) ≟ lit (true ∷ []) + then lit ℤ.1ℤ + else lit ℤ.0ℤ) + +!_ : Reference Σ Γ t → Expression Σ Γ t +! state i = state i +! var i = var i +! [ ref ] = [ ! ref ] +! unbox ref = unbox (! ref) +! merge ref ref₁ e = merge (! ref) (! ref₁) e +! slice ref x = slice (! ref) x +! cut ref x = cut (! ref) x +! cast eq ref = cast eq (! ref) +! nil = nil +! cons ref ref₁ = cons (! ref) (! ref₁) +! head ref = head (! ref) +! tail ref = tail (! ref) + +!!_ : LocalReference Σ Γ t → Expression Σ Γ t +!! var i = var i +!! [ ref ] = [ !! ref ] +!! unbox ref = unbox (!! ref) +!! merge ref ref₁ x = merge (!! ref) (!! ref₁) x +!! slice ref x = slice (!! ref) x +!! cut ref x = cut (!! ref) x +!! cast eq ref = cast eq (!! ref) +!! nil = nil +!! cons ref ref₁ = cons (!! ref) (!! ref₁) +!! head ref = head (!! ref) +!! tail ref = tail (!! ref) diff --git a/src/Helium/Data/Pseudocode/Manipulate.agda b/src/Helium/Data/Pseudocode/Manipulate.agda index a798ad8..d37cfc9 100644 --- a/src/Helium/Data/Pseudocode/Manipulate.agda +++ b/src/Helium/Data/Pseudocode/Manipulate.agda @@ -6,1549 +6,1271 @@ {-# OPTIONS --safe --without-K #-} -open import Data.Vec using (Vec) -open import Helium.Data.Pseudocode.Core +module Helium.Data.Pseudocode.Manipulate where -module Helium.Data.Pseudocode.Manipulate - {o} {Σ : Vec Type o} - where +open import Helium.Data.Pseudocode.Core -import Algebra.Solver.IdempotentCommutativeMonoid as ComMonoidSolver -open import Data.Fin as Fin using (Fin; suc) +open import Algebra.Bundles using (IdempotentCommutativeMonoid) +import Algebra.Solver.IdempotentCommutativeMonoid as ICMSolver +import Algebra.Solver.Ring.AlmostCommutativeRing as ACR +import Algebra.Solver.Ring.Simple as RingSolver +open import Data.Fin as Fin using (suc; punchOut; punchIn) open import Data.Fin.Patterns -open import Data.Nat as ℕ using (ℕ; suc; _⊔_) +open import Data.Nat as ℕ using (ℕ; suc; _<_; _≤_; z≤n; s≤s; _⊔_) import Data.Nat.Induction as ℕᵢ import Data.Nat.Properties as ℕₚ -open import Data.Nat.Solver using (module +-*-Solver) -open import Data.Product using (∃; _×_; _,_; proj₁; proj₂; <_,_>) -open import Data.Sum using (_⊎_; inj₁; inj₂) -import Data.Product.Relation.Binary.Lex.Strict as Lex -open import Data.Vec as Vec using ([]; _∷_) +open import Data.Product using (∃; _×_; _,_; proj₁; proj₂; -,_; <_,_>) +open import Data.Product.Nary.NonDependent using (Product; uncurryₙ) +open import Data.Product.Relation.Binary.Lex.Strict +open import Data.Sum using (inj₁; inj₂) +open import Data.Unit.Polymorphic using (⊤) +open import Data.Vec as Vec using (Vec; []; _∷_; lookup; insert; remove) import Data.Vec.Properties as Vecₚ +open import Data.Vec.Recursive as Vecᵣ using (2+_) open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -open import Data.Vec.Relation.Unary.Any using (Any; here; there) -open import Function using (_on_; _∘_; _∋_; case_return_of_) -open import Function.Nary.NonDependent using (congₙ) -open import Helium.Data.Pseudocode.Properties -import Induction.WellFounded as Wf +open import Function +open import Function.Nary.NonDependent using (_⇉_; Sets; congₙ; 0ℓs) +open import Helium.Data.Pseudocode.Core +open import Induction.WellFounded as Wf using (WellFounded) import Relation.Binary.Construct.On as On -open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; cong; cong₂; module ≡-Reasoning) -open import Relation.Nullary using (yes; no; ¬_) -open import Relation.Nullary.Decidable.Core using (True; fromWitness; toWitness; toWitnessFalse) -open import Relation.Nullary.Negation using (contradiction) +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary using (yes; no) + +private + variable + k m n o : ℕ + ret t t′ t₁ t₂ : Type + Σ Γ Δ ts : Vec Type n -open ComMonoidSolver (record +private + punchOut-insert : ∀ {a} {A : Set a} (xs : Vec A n) {i j} (i≢j : i ≢ j) x → lookup xs (punchOut i≢j) ≡ lookup (insert xs i x) j + punchOut-insert xs {i} {j} i≢j x = begin + lookup xs (punchOut i≢j) ≡˘⟨ cong (flip lookup (punchOut i≢j)) (Vecₚ.remove-insert xs i x) ⟩ + lookup (remove (insert xs i x) i) (punchOut i≢j) ≡⟨ Vecₚ.remove-punchOut (insert xs i x) i≢j ⟩ + lookup (insert xs i x) j ∎ + where open ≡-Reasoning + + lookupᵣ-map : ∀ {a b} {A : Set a} {B : Set b} {f : A → B} i (xs : A Vecᵣ.^ n) → Vecᵣ.lookup i (Vecᵣ.map f n xs) ≡ f (Vecᵣ.lookup i xs) + lookupᵣ-map {n = 1} 0F xs = refl + lookupᵣ-map {n = 2+ n} 0F xs = refl + lookupᵣ-map {n = 2+ n} (suc i) xs = lookupᵣ-map i (proj₂ xs) + + ⊔-0-boundedLattice : IdempotentCommutativeMonoid _ _ + ⊔-0-boundedLattice = record { isIdempotentCommutativeMonoid = record { isCommutativeMonoid = ℕₚ.⊔-0-isCommutativeMonoid - ; idem = ℕₚ.⊔-idem + ; idem = ℕₚ.⊔-idem } - }) - using (_⊕_; _⊜_) - renaming (solve to ⊔-solve) + } -open Code Σ +open ICMSolver ⊔-0-boundedLattice + using (_⊜_; _⊕_) + renaming (solve to solve-⊔; id to ε) -private - variable - m n : ℕ - Γ Δ : Vec Type m - t t′ ret : Type - --- TODO: make argument irrelevant -castType : Expression Γ t → (t ≡ t′) → Expression Γ t′ -castType e refl = e - -cast-pres-assignable : ∀ {e : Expression Γ t} → CanAssign e → (eq : t ≡ t′) → CanAssign (castType e eq) -cast-pres-assignable e refl = e - -cast-pres-stateless : ∀ {e : Expression Γ t} → (eq : t ≡ t′) → ReferencesState (castType e eq) → ReferencesState e -cast-pres-stateless refl e = e - -punchOut⇒insert : ∀ {a} {A : Set a} (xs : Vec A n) {i j : Fin (suc n)} (i≢j : ¬ i ≡ j) x → Vec.lookup xs (Fin.punchOut i≢j) ≡ Vec.lookup (Vec.insert xs i x) j -punchOut⇒insert xs {i} {j} i≢j x = begin - Vec.lookup xs (Fin.punchOut i≢j) - ≡˘⟨ cong (λ x → Vec.lookup x _) (Vecₚ.remove-insert xs i x) ⟩ - Vec.lookup (Vec.remove (Vec.insert xs i x) i) (Fin.punchOut i≢j) - ≡⟨ Vecₚ.remove-punchOut (Vec.insert xs i x) i≢j ⟩ - Vec.lookup (Vec.insert xs i x) j - ∎ - where open ≡-Reasoning - -elimAt : ∀ i → Expression (Vec.insert Γ i t′) t → Expression Γ t′ → Expression Γ t -elimAt′ : ∀ i → All (Expression (Vec.insert Γ i t′)) Δ → Expression Γ t′ → All (Expression Γ) Δ - -elimAt i (lit x) e′ = lit x -elimAt i (state j) e′ = state j -elimAt i (var j) e′ with i Fin.≟ j -... | yes refl = castType e′ (sym (Vecₚ.insert-lookup _ i _)) -... | no i≢j = castType (var (Fin.punchOut i≢j)) (punchOut⇒insert _ i≢j _) -elimAt i (abort e) e′ = abort (elimAt i e e′) -elimAt i (_≟_ {hasEquality = hasEq} e e₁) e′ = _≟_ {hasEquality = hasEq} (elimAt i e e′) (elimAt i e₁ e′) -elimAt i (_<?_ {isNumeric = isNum} e e₁) e′ = _<?_ {isNumeric = isNum} (elimAt i e e′) (elimAt i e₁ e′) -elimAt i (inv e) e′ = elimAt i e e′ -elimAt i (e && e₁) e′ = elimAt i e e′ && elimAt i e₁ e′ -elimAt i (e || e₁) e′ = elimAt i e e′ || elimAt i e₁ e′ -elimAt i (not e) e′ = not (elimAt i e e′) -elimAt i (e and e₁) e′ = elimAt i e e′ and elimAt i e₁ e′ -elimAt i (e or e₁) e′ = elimAt i e e′ or elimAt i e₁ e′ -elimAt i [ e ] e′ = [ elimAt i e e′ ] -elimAt i (unbox e) e′ = unbox (elimAt i e e′) -elimAt i (splice e e₁ e₂) e′ = splice (elimAt i e e′) (elimAt i e₁ e′) (elimAt i e₂ e′) -elimAt i (cut e e₁) e′ = cut (elimAt i e e′) (elimAt i e₁ e′) -elimAt i (cast eq e) e′ = cast eq (elimAt i e e′) -elimAt i (-_ {isNumeric = isNum} e) e′ = -_ {isNumeric = isNum} (elimAt i e e′) -elimAt i (e + e₁) e′ = elimAt i e e′ + elimAt i e₁ e′ -elimAt i (e * e₁) e′ = elimAt i e e′ * elimAt i e₁ e′ -elimAt i (_^_ {isNumeric = isNum} e x) e′ = _^_ {isNumeric = isNum} (elimAt i e e′) x -elimAt i (e >> x) e′ = elimAt i e e′ >> x -elimAt i (rnd e) e′ = rnd (elimAt i e e′) -elimAt i (fin x e) e′ = fin x (elimAt i e e′) -elimAt i (asInt e) e′ = asInt (elimAt i e e′) -elimAt i nil e′ = nil -elimAt i (cons e e₁) e′ = cons (elimAt i e e′) (elimAt i e₁ e′) -elimAt i (head e) e′ = head (elimAt i e e′) -elimAt i (tail e) e′ = tail (elimAt i e e′) -elimAt i (call f es) e′ = call f (elimAt′ i es e′) -elimAt i (if e then e₁ else e₂) e′ = if elimAt i e e′ then elimAt i e₁ e′ else elimAt i e₂ e′ - -elimAt′ i [] e′ = [] -elimAt′ i (e ∷ es) e′ = elimAt i e e′ ∷ elimAt′ i es e′ - -wknAt : ∀ i → Expression Γ t → Expression (Vec.insert Γ i t′) t -wknAt′ : ∀ i → All (Expression Γ) Δ → All (Expression (Vec.insert Γ i t′)) Δ - -wknAt i (lit x) = lit x -wknAt i (state j) = state j -wknAt i (var j) = castType (var (Fin.punchIn i j)) (Vecₚ.insert-punchIn _ i _ j) -wknAt i (abort e) = abort (wknAt i e) -wknAt i (_≟_ {hasEquality = hasEq} e e₁) = _≟_ {hasEquality = hasEq} (wknAt i e) (wknAt i e₁) -wknAt i (_<?_ {isNumeric = isNum} e e₁) = _<?_ {isNumeric = isNum} (wknAt i e) (wknAt i e₁) -wknAt i (inv e) = inv (wknAt i e) -wknAt i (e && e₁) = wknAt i e && wknAt i e₁ -wknAt i (e || e₁) = wknAt i e && wknAt i e₁ -wknAt i (not e) = not (wknAt i e) -wknAt i (e and e₁) = wknAt i e and wknAt i e₁ -wknAt i (e or e₁) = wknAt i e or wknAt i e₁ -wknAt i [ e ] = [ wknAt i e ] -wknAt i (unbox e) = unbox (wknAt i e) -wknAt i (splice e e₁ e₂) = splice (wknAt i e) (wknAt i e₁) (wknAt i e₂) -wknAt i (cut e e₁) = cut (wknAt i e) (wknAt i e₁) -wknAt i (cast eq e) = cast eq (wknAt i e) -wknAt i (-_ {isNumeric = isNum} e) = -_ {isNumeric = isNum} (wknAt i e) -wknAt i (e + e₁) = wknAt i e + wknAt i e₁ -wknAt i (e * e₁) = wknAt i e * wknAt i e₁ -wknAt i (_^_ {isNumeric = isNum} e x) = _^_ {isNumeric = isNum} (wknAt i e) x -wknAt i (e >> x) = wknAt i e >> x -wknAt i (rnd e) = rnd (wknAt i e) -wknAt i (fin x e) = fin x (wknAt i e) -wknAt i (asInt e) = asInt (wknAt i e) -wknAt i nil = nil -wknAt i (cons e e₁) = cons (wknAt i e) (wknAt i e₁) -wknAt i (head e) = head (wknAt i e) -wknAt i (tail e) = tail (wknAt i e) -wknAt i (call f es) = call f (wknAt′ i es) -wknAt i (if e then e₁ else e₂) = if wknAt i e then wknAt i e₁ else wknAt i e₂ - -wknAt′ i [] = [] -wknAt′ i (e ∷ es) = wknAt i e ∷ wknAt′ i es - -substAt : ∀ i → Expression Γ t → Expression Γ (Vec.lookup Γ i) → Expression Γ t -substAt′ : ∀ i → All (Expression Γ) Δ → Expression Γ (Vec.lookup Γ i) → All (Expression Γ) Δ -substAt i (lit x) e′ = lit x -substAt i (state j) e′ = state j -substAt i (var j) e′ with i Fin.≟ j -... | yes refl = e′ -... | no _ = var j -substAt i (abort e) e′ = abort (substAt i e e′) -substAt i (_≟_ {hasEquality = hasEq} e e₁) e′ = _≟_ {hasEquality = hasEq} (substAt i e e′) (substAt i e₁ e′) -substAt i (_<?_ {isNumeric = isNum} e e₁) e′ = _<?_ {isNumeric = isNum} (substAt i e e′) (substAt i e₁ e′) -substAt i (inv e) e′ = inv (substAt i e e′) -substAt i (e && e₁) e′ = substAt i e e′ && substAt i e₁ e′ -substAt i (e || e₁) e′ = substAt i e e′ || substAt i e₁ e′ -substAt i (not e) e′ = not (substAt i e e′) -substAt i (e and e₁) e′ = substAt i e e′ and substAt i e₁ e′ -substAt i (e or e₁) e′ = substAt i e e′ or substAt i e₁ e′ -substAt i [ e ] e′ = [ substAt i e e′ ] -substAt i (unbox e) e′ = unbox (substAt i e e′) -substAt i (splice e e₁ e₂) e′ = splice (substAt i e e′) (substAt i e₁ e′) (substAt i e₂ e′) -substAt i (cut e e₁) e′ = cut (substAt i e e′) (substAt i e₁ e′) -substAt i (cast eq e) e′ = cast eq (substAt i e e′) -substAt i (-_ {isNumeric = isNum} e) e′ = -_ {isNumeric = isNum} (substAt i e e′) -substAt i (e + e₁) e′ = substAt i e e′ + substAt i e₁ e′ -substAt i (e * e₁) e′ = substAt i e e′ * substAt i e₁ e′ -substAt i (_^_ {isNumeric = isNum} e x) e′ = _^_ {isNumeric = isNum} (substAt i e e′) x -substAt i (e >> x) e′ = substAt i e e′ >> x -substAt i (rnd e) e′ = rnd (substAt i e e′) -substAt i (fin x e) e′ = fin x (substAt i e e′) -substAt i (asInt e) e′ = asInt (substAt i e e′) -substAt i nil e′ = nil -substAt i (cons e e₁) e′ = cons (substAt i e e′) (substAt i e₁ e′) -substAt i (head e) e′ = head (substAt i e e′) -substAt i (tail e) e′ = tail (substAt i e e′) -substAt i (call f es) e′ = call f (substAt′ i es e′) -substAt i (if e then e₁ else e₂) e′ = if substAt i e e′ then substAt i e₁ e′ else substAt i e₂ e′ - -substAt′ i [] e′ = [] -substAt′ i (e ∷ es) e′ = substAt i e e′ ∷ substAt′ i es e′ - -updateRef : ∀ {e : Expression Γ t} (ref : CanAssign e) → ¬ ReferencesState e → Expression Γ t → Expression Γ t′ → Expression Γ t′ -updateRef (state i) stateless val e = contradiction (state i) stateless -updateRef (var i) stateless val e = substAt i e val -updateRef (abort _) stateless val e = e -updateRef [ ref ] stateless val e = updateRef ref (stateless ∘ [_]) (unbox val) e -updateRef (unbox ref) stateless val e = updateRef ref (stateless ∘ unbox) [ val ] e -updateRef (splice ref ref₁ e₂) stateless val e = updateRef ref₁ (stateless ∘ (λ x → spliceʳ _ x _)) (head (tail (cut val e₂))) (updateRef ref (stateless ∘ (λ x → spliceˡ x _ _)) (head (cut val e₂)) e) -updateRef (cut ref e₁) stateless val e = updateRef ref (stateless ∘ (λ x → cut x _)) (splice (head val) (head (tail val)) e₁) e -updateRef (cast eq ref) stateless val e = updateRef ref (stateless ∘ cast eq) (cast (sym eq) val) e -updateRef nil stateless val e = e -updateRef (cons ref ref₁) stateless val e = updateRef ref₁ (stateless ∘ (λ x → consʳ _ x)) (tail val) (updateRef ref (stateless ∘ (λ x → consˡ x _)) (head val) e) -updateRef (head {e = e′} ref) stateless val e = updateRef ref (stateless ∘ head) (cons val (tail e′)) e -updateRef (tail {e = e′} ref) stateless val e = updateRef ref (stateless ∘ tail) (cons (head e′) val) e - -wknAt-pres-assignable : ∀ i {e} → CanAssign e → CanAssign (wknAt {Γ = Γ} {t} {t′} i e) -wknAt-pres-assignable i (state j) = state j -wknAt-pres-assignable i (var j) = cast-pres-assignable (var (Fin.punchIn i j)) (Vecₚ.insert-punchIn _ i _ j) -wknAt-pres-assignable i (abort e) = abort (wknAt i e) -wknAt-pres-assignable i [ ref ] = [ wknAt-pres-assignable i ref ] -wknAt-pres-assignable i (unbox ref) = unbox (wknAt-pres-assignable i ref) -wknAt-pres-assignable i (splice ref ref₁ e₂) = splice (wknAt-pres-assignable i ref) (wknAt-pres-assignable i ref₁) (wknAt i e₂) -wknAt-pres-assignable i (cut ref e₁) = cut (wknAt-pres-assignable i ref) (wknAt i e₁) -wknAt-pres-assignable i (cast eq ref) = cast eq (wknAt-pres-assignable i ref) -wknAt-pres-assignable i nil = nil -wknAt-pres-assignable i (cons ref ref₁) = cons (wknAt-pres-assignable i ref) (wknAt-pres-assignable i ref₁) -wknAt-pres-assignable i (head ref) = head (wknAt-pres-assignable i ref) -wknAt-pres-assignable i (tail ref) = tail (wknAt-pres-assignable i ref) - -wknAt-pres-stateless : ∀ i {e} → ReferencesState (wknAt {Γ = Γ} {t} {t′} i e) → ReferencesState e -wknAt-pres-stateless i {state _} (state j) = state j -wknAt-pres-stateless i {var j} e = contradiction (cast-pres-stateless {e = var (Fin.punchIn i j)} (Vecₚ.insert-punchIn _ i _ j) e) (λ ()) -wknAt-pres-stateless i {[ _ ]} [ e ] = [ wknAt-pres-stateless i e ] -wknAt-pres-stateless i {unbox _} (unbox e) = unbox (wknAt-pres-stateless i e) -wknAt-pres-stateless i {splice _ _ _} (spliceˡ e e₁ e₂) = spliceˡ (wknAt-pres-stateless i e) _ _ -wknAt-pres-stateless i {splice _ _ _} (spliceʳ e e₁ e₂) = spliceʳ _ (wknAt-pres-stateless i e₁) _ -wknAt-pres-stateless i {cut _ _} (cut e e₁) = cut (wknAt-pres-stateless i e) _ -wknAt-pres-stateless i {cast _ _} (cast eq e) = cast eq (wknAt-pres-stateless i e) -wknAt-pres-stateless i {cons _ _} (consˡ e e₁) = consˡ (wknAt-pres-stateless i e) _ -wknAt-pres-stateless i {cons _ _} (consʳ e e₁) = consʳ _ (wknAt-pres-stateless i e₁) -wknAt-pres-stateless i {head _} (head e) = head (wknAt-pres-stateless i e) -wknAt-pres-stateless i {tail _} (tail e) = tail (wknAt-pres-stateless i e) - -wknStatementAt : ∀ t i → Statement Γ → Statement (Vec.insert Γ i t) -wknStatementAt t i (s ∙ s₁) = wknStatementAt t i s ∙ wknStatementAt t i s₁ -wknStatementAt t i skip = skip -wknStatementAt t i (_≔_ ref {assignable} x) = _≔_ (wknAt i ref) {fromWitness (wknAt-pres-assignable i (toWitness assignable))} (wknAt i x) -wknStatementAt t i (declare x s) = declare (wknAt i x) (wknStatementAt t (suc i) s) -wknStatementAt t i (invoke p es) = invoke p (wknAt′ i es) -wknStatementAt t i (if x then s) = if wknAt i x then wknStatementAt t i s -wknStatementAt t i (if x then s else s₁) = if wknAt i x then wknStatementAt t i s else wknStatementAt t i s₁ -wknStatementAt t i (for m s) = for m (wknStatementAt t (suc i) s) - -subst : Expression Γ t → All (Expression Δ) Γ → Expression Δ t -subst′ : ∀ {n ts} → All (Expression Γ) {n} ts → All (Expression Δ) Γ → All (Expression Δ) ts - -subst (lit x) xs = lit x -subst (state i) xs = state i -subst (var i) xs = All.lookup i xs -subst (abort e) xs = abort (subst e xs) -subst (_≟_ {hasEquality = hasEq} e e₁) xs = _≟_ {hasEquality = hasEq} (subst e xs) (subst e₁ xs) -subst (_<?_ {isNumeric = isNum} e e₁) xs = _<?_ {isNumeric = isNum} (subst e xs) (subst e₁ xs) -subst (inv e) xs = inv (subst e xs) -subst (e && e₁) xs = subst e xs && subst e₁ xs -subst (e || e₁) xs = subst e xs || subst e₁ xs -subst (not e) xs = not (subst e xs) -subst (e and e₁) xs = subst e xs and subst e₁ xs -subst (e or e₁) xs = subst e xs or subst e₁ xs -subst [ e ] xs = [ subst e xs ] -subst (unbox e) xs = unbox (subst e xs) -subst (splice e e₁ e₂) xs = splice (subst e xs) (subst e₁ xs) (subst e₂ xs) -subst (cut e e₁) xs = cut (subst e xs) (subst e₁ xs) -subst (cast eq e) xs = cast eq (subst e xs) -subst (-_ {isNumeric = isNum} e) xs = -_ {isNumeric = isNum} (subst e xs) -subst (e + e₁) xs = subst e xs + subst e₁ xs -subst (e * e₁) xs = subst e xs * subst e₁ xs -subst (_^_ {isNumeric = isNum} e x) xs = _^_ {isNumeric = isNum} (subst e xs) x -subst (e >> x) xs = subst e xs >> x -subst (rnd e) xs = rnd (subst e xs) -subst (fin x e) xs = fin x (subst e xs) -subst (asInt e) xs = asInt (subst e xs) -subst nil xs = nil -subst (cons e e₁) xs = cons (subst e xs) (subst e₁ xs) -subst (head e) xs = head (subst e xs) -subst (tail e) xs = tail (subst e xs) -subst (call f es) xs = call f (subst′ es xs) -subst (if e then e₁ else e₂) xs = if subst e xs then subst e₁ xs else subst e₂ xs - -subst′ [] xs = [] -subst′ (e ∷ es) xs = subst e xs ∷ subst′ es xs - -callDepth : Expression Γ t → ℕ -callDepth′ : All (Expression Γ) Δ → ℕ -stmtCallDepth : Statement Γ → ℕ -funCallDepth : Function Γ ret → ℕ -procCallDepth : Procedure Γ → ℕ - -callDepth (lit x) = 0 -callDepth (state i) = 0 -callDepth (var i) = 0 -callDepth (abort e) = callDepth e -callDepth (e ≟ e₁) = callDepth e ⊔ callDepth e₁ -callDepth (e <? e₁) = callDepth e ⊔ callDepth e₁ -callDepth (inv e) = callDepth e -callDepth (e && e₁) = callDepth e ⊔ callDepth e₁ -callDepth (e || e₁) = callDepth e ⊔ callDepth e₁ -callDepth (not e) = callDepth e -callDepth (e and e₁) = callDepth e ⊔ callDepth e₁ -callDepth (e or e₁) = callDepth e ⊔ callDepth e₁ -callDepth [ e ] = callDepth e -callDepth (unbox e) = callDepth e -callDepth (splice e e₁ e₂) = callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂ -callDepth (cut e e₁) = callDepth e ⊔ callDepth e₁ -callDepth (cast eq e) = callDepth e -callDepth (- e) = callDepth e -callDepth (e + e₁) = callDepth e ⊔ callDepth e₁ -callDepth (e * e₁) = callDepth e ⊔ callDepth e₁ -callDepth (e ^ x) = callDepth e -callDepth (e >> x) = callDepth e -callDepth (rnd e) = callDepth e -callDepth (fin x e) = callDepth e -callDepth (asInt e) = callDepth e -callDepth nil = 0 -callDepth (cons e e₁) = callDepth e ⊔ callDepth e₁ -callDepth (head e) = callDepth e -callDepth (tail e) = callDepth e -callDepth (call f es) = suc (funCallDepth f) ⊔ callDepth′ es -callDepth (if e then e₁ else e₂) = callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂ - -callDepth′ [] = 0 -callDepth′ (e ∷ es) = callDepth e ⊔ callDepth′ es - -stmtCallDepth (s ∙ s₁) = stmtCallDepth s ⊔ stmtCallDepth s₁ -stmtCallDepth skip = 0 -stmtCallDepth (ref ≔ x) = callDepth ref ⊔ callDepth x -stmtCallDepth (declare x s) = callDepth x ⊔ stmtCallDepth s -stmtCallDepth (invoke p es) = procCallDepth p ⊔ callDepth′ es -stmtCallDepth (if x then s) = callDepth x ⊔ stmtCallDepth s -stmtCallDepth (if x then s else s₁) = callDepth x ⊔ stmtCallDepth s ⊔ stmtCallDepth s₁ -stmtCallDepth (for m s) = stmtCallDepth s - -funCallDepth (s ∙return x) = stmtCallDepth s ⊔ callDepth x -funCallDepth (declare x f) = funCallDepth f ⊔ callDepth x - -procCallDepth (x ∙end) = stmtCallDepth x +open RingSolver (ACR.fromCommutativeSemiring ℕₚ.+-*-commutativeSemiring) ℕₚ._≟_ + using (_:=_; _:+_; _:*_; con) + renaming (solve to solve-+) open ℕₚ.≤-Reasoning -castType-pres-callDepth : ∀ (e : Expression Γ t) (eq : t ≡ t′) → callDepth (castType e eq) ≡ callDepth e -castType-pres-callDepth e refl = refl - -elimAt-pres-callDepth : ∀ i (e : Expression (Vec.insert Γ i t′) t) (e′ : Expression Γ t′) → callDepth (elimAt i e e′) ℕ.≤ callDepth e′ ⊔ callDepth e -elimAt′-pres-callDepth : ∀ i (es : All (Expression (Vec.insert Γ i t′)) Δ) (e′ : Expression Γ t′) → callDepth′ (elimAt′ i es e′) ℕ.≤ callDepth e′ ⊔ callDepth′ es - -elimAt-pres-callDepth i (lit x) e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -elimAt-pres-callDepth i (state j) e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -elimAt-pres-callDepth i (var j) e′ with i Fin.≟ j -... | yes refl = begin - callDepth (castType e′ (sym (Vecₚ.insert-lookup _ i _))) - ≡⟨ castType-pres-callDepth e′ (sym (Vecₚ.insert-lookup _ i _)) ⟩ - callDepth e′ - ≤⟨ ℕₚ.m≤m⊔n (callDepth e′) 0 ⟩ - callDepth e′ ⊔ 0 - ∎ -elimAt-pres-callDepth {Γ = Γ} i (var j) e′ | no i≢j = begin - callDepth (castType (var {Γ = Γ} (Fin.punchOut i≢j)) (punchOut⇒insert Γ i≢j _)) - ≡⟨ castType-pres-callDepth (var {Γ = Γ} (Fin.punchOut i≢j)) (punchOut⇒insert Γ i≢j _) ⟩ - 0 - ≤⟨ ℕ.z≤n ⟩ - callDepth e′ ⊔ 0 +private + [_]_$_⊗_ : ∀ {a b c} {A : Set a} {B : Set b} m → (C : A → B → Set c) → A Vecᵣ.^ m → B Vecᵣ.^ m → Set c + [ m ] f $ xs ⊗ ys = Vecᵣ.foldr ⊤ id (const _×_) m (Vecᵣ.zipWith f m xs ys) + + ⨆[_]_ : ∀ n → ℕ Vecᵣ.^ n → ℕ + ⨆[_]_ = Vecᵣ.foldl (const ℕ) 0 id (const (flip _⊔_)) + + ⨆-step : ∀ m x xs → ⨆[ 2+ m ] (x , xs) ≡ x ⊔ ⨆[ suc m ] xs + ⨆-step 0 x xs = refl + ⨆-step (suc m) x (y , xs) = begin-equality + ⨆[ 2+ suc m ] (x , y , xs) ≡⟨ ⨆-step m (x ⊔ y) xs ⟩ + x ⊔ y ⊔ ⨆[ suc m ] xs ≡⟨ ℕₚ.⊔-assoc x y _ ⟩ + x ⊔ (y ⊔ ⨆[ suc m ] xs) ≡˘⟨ cong (_ ⊔_) (⨆-step m y xs) ⟩ + x ⊔ ⨆[ 2+ m ] (y , xs) ∎ + + join-lubs : ∀ n m {lhs rhs} → [ m ] (λ x y → x ≤ y ⊔ n) $ lhs ⊗ rhs → ⨆[ m ] lhs ≤ (⨆[ m ] rhs) ⊔ n + join-lubs n 0 {lhs} {rhs} ≤s = z≤n + join-lubs n 1 {lhs} {rhs} ≤s = ≤s + join-lubs n (2+ m) {x , lhs} {y , rhs} (x≤y⊔n , ≤s) = begin + ⨆[ 2+ m ] (x , lhs) ≡⟨ ⨆-step m x lhs ⟩ + x ⊔ ⨆[ suc m ] lhs ≤⟨ ℕₚ.⊔-mono-≤ x≤y⊔n (join-lubs n (suc m) ≤s) ⟩ + y ⊔ n ⊔ (⨆[ suc m ] rhs ⊔ n) ≡⟨ solve-⊔ 3 (λ a b c → (a ⊕ c) ⊕ b ⊕ c ⊜ (a ⊕ b) ⊕ c) refl y _ n ⟩ + y ⊔ ⨆[ suc m ] rhs ⊔ n ≡˘⟨ cong (_⊔ _) (⨆-step m y rhs) ⟩ + ⨆[ 2+ m ] (y , rhs) ⊔ n ∎ + + lookup-⨆-≤ : ∀ i (xs : ℕ Vecᵣ.^ n) → Vecᵣ.lookup i xs ≤ ⨆[ n ] xs + lookup-⨆-≤ {1} 0F x = ℕₚ.≤-refl + lookup-⨆-≤ {2+ n} 0F (x , xs) = begin + x ≤⟨ ℕₚ.m≤m⊔n x _ ⟩ + x ⊔ ⨆[ suc n ] xs ≡˘⟨ ⨆-step n x xs ⟩ + ⨆[ 2+ n ] (x , xs) ∎ + lookup-⨆-≤ {2+ n} (suc i) (x , xs) = begin + Vecᵣ.lookup i xs ≤⟨ lookup-⨆-≤ i xs ⟩ + ⨆[ suc n ] xs ≤⟨ ℕₚ.m≤n⊔m x _ ⟩ + x ⊔ ⨆[ suc n ] xs ≡˘⟨ ⨆-step n x xs ⟩ + ⨆[ 2+ n ] (x , xs) ∎ + + Σ[_]_ : ∀ n → ℕ Vecᵣ.^ n → ℕ + Σ[_]_ = Vecᵣ.foldl (const ℕ) 0 id (const (flip ℕ._+_)) + + Σ-step : ∀ m x xs → Σ[ 2+ m ] (x , xs) ≡ x ℕ.+ Σ[ suc m ] xs + Σ-step 0 x xs = refl + Σ-step (suc m) x (y , xs) = begin-equality + Σ[ 2+ suc m ] (x , y , xs) ≡⟨ Σ-step m (x ℕ.+ y) xs ⟩ + x ℕ.+ y ℕ.+ Σ[ suc m ] xs ≡⟨ ℕₚ.+-assoc x y _ ⟩ + x ℕ.+ (y ℕ.+ Σ[ suc m ] xs) ≡˘⟨ cong (x ℕ.+_) (Σ-step m y xs) ⟩ + x ℕ.+ Σ[ 2+ m ] (y , xs) ∎ + + lookup-Σ-≤ : ∀ i (xs : ℕ Vecᵣ.^ n) → Vecᵣ.lookup i xs ≤ Σ[ n ] xs + lookup-Σ-≤ {1} 0F x = ℕₚ.≤-refl + lookup-Σ-≤ {2+ n} 0F (x , xs) = begin + x ≤⟨ ℕₚ.m≤m+n x _ ⟩ + x ℕ.+ Σ[ suc n ] xs ≡˘⟨ Σ-step n x xs ⟩ + Σ[ 2+ n ] (x , xs) ∎ + lookup-Σ-≤ {2+ n} (suc i) (x , xs) = begin + Vecᵣ.lookup i xs ≤⟨ lookup-Σ-≤ i xs ⟩ + Σ[ suc n ] xs ≤⟨ ℕₚ.m≤n+m _ x ⟩ + x ℕ.+ Σ[ suc n ] xs ≡˘⟨ Σ-step n x xs ⟩ + Σ[ 2+ n ] (x , xs) ∎ + + + foldr-lubs : ∀ {a b c} {A : Set a} {B : ℕ → Set b} + (f : ∀ {n} → A → B n → B (suc n)) y + (P : ∀ {n} → B n → Set c) → + P y → + (∀ {n} a {b : B n} → P b → P (f a b)) → + ∀ (xs : Vec A n) → + P (Vec.foldr B f y xs) + foldr-lubs f y P y∈P f-pres [] = y∈P + foldr-lubs f y P y∈P f-pres (x ∷ xs) = f-pres x (foldr-lubs f y P y∈P f-pres xs) + +module CallDepth where + expr : Expression Σ Γ t → ℕ + exprs : All (Expression Σ Γ) ts → ℕ + locRef : LocalReference Σ Γ t → ℕ + locStmt : LocalStatement Σ Γ → ℕ + fun : Function Σ Γ ret → ℕ + + expr (lit x) = 0 + expr (state i) = 0 + expr (var i) = 0 + expr (e ≟ e₁) = expr e ⊔ expr e₁ + expr (e <? e₁) = expr e ⊔ expr e₁ + expr (inv e) = expr e + expr (e && e₁) = expr e ⊔ expr e₁ + expr (e || e₁) = expr e ⊔ expr e₁ + expr (not e) = expr e + expr (e and e₁) = expr e ⊔ expr e₁ + expr (e or e₁) = expr e ⊔ expr e₁ + expr [ e ] = expr e + expr (unbox e) = expr e + expr (merge e e₁ e₂) = expr e ⊔ expr e₁ ⊔ expr e₂ + expr (slice e e₁) = expr e ⊔ expr e₁ + expr (cut e e₁) = expr e ⊔ expr e₁ + expr (cast eq e) = expr e + expr (- e) = expr e + expr (e + e₁) = expr e ⊔ expr e₁ + expr (e * e₁) = expr e ⊔ expr e₁ + expr (e ^ x) = expr e + expr (e >> n) = expr e + expr (rnd e) = expr e + expr (fin f e) = expr e + expr (asInt e) = expr e + expr nil = 0 + expr (cons e e₁) = expr e ⊔ expr e₁ + expr (head e) = expr e + expr (tail e) = expr e + expr (call f es) = exprs es ⊔ suc (fun f) + expr (if e then e₁ else e₂) = expr e ⊔ expr e₁ ⊔ expr e₂ + + exprs [] = 0 + exprs (e ∷ es) = exprs es ⊔ expr e + + locRef (var i) = 0 + locRef [ ref ] = locRef ref + locRef (unbox ref) = locRef ref + locRef (merge ref ref₁ x) = locRef ref ⊔ locRef ref₁ ⊔ expr x + locRef (slice ref x) = locRef ref ⊔ expr x + locRef (cut ref x) = locRef ref ⊔ expr x + locRef (cast eq ref) = locRef ref + locRef nil = 0 + locRef (cons ref ref₁) = locRef ref ⊔ locRef ref₁ + locRef (head ref) = locRef ref + locRef (tail ref) = locRef ref + + locStmt (s ∙ s₁) = locStmt s ⊔ locStmt s₁ + locStmt skip = 0 + locStmt (ref ≔ e) = locRef ref ⊔ expr e + locStmt (declare x s) = locStmt s ⊔ expr x + locStmt (if x then s) = locStmt s ⊔ expr x + locStmt (if x then s else s₁) = locStmt s ⊔ locStmt s₁ ⊔ expr x + locStmt (for n s) = locStmt s + + fun (declare x f) = fun f ⊔ expr x + fun (s ∙return e) = locStmt s ⊔ expr e + + homo-!! : ∀ (ref : LocalReference Σ Γ t) → expr (!! ref) ≡ locRef ref + homo-!! (var i) = refl + homo-!! [ ref ] = homo-!! ref + homo-!! (unbox ref) = homo-!! ref + homo-!! (merge ref ref₁ x) = cong₂ (λ m n → m ⊔ n ⊔ _) (homo-!! ref) (homo-!! ref₁) + homo-!! (slice ref x) = cong (_⊔ _) (homo-!! ref) + homo-!! (cut ref x) = cong (_⊔ _) (homo-!! ref) + homo-!! (cast eq ref) = homo-!! ref + homo-!! nil = refl + homo-!! (cons ref ref₁) = cong₂ _⊔_ (homo-!! ref) (homo-!! ref₁) + homo-!! (head ref) = homo-!! ref + homo-!! (tail ref) = homo-!! ref + + ∷ˡ-≤ : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → + expr e ≤ exprs (e ∷ es) + ∷ˡ-≤ e es = ℕₚ.m≤n⊔m (exprs es) (expr e) + + ∷ʳ-≤ : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → + exprs es ≤ exprs (e ∷ es) + ∷ʳ-≤ e es = ℕₚ.m≤m⊔n (exprs es) (expr e) + + lookup-≤ : ∀ i (es : All (Expression Σ Γ) ts) → expr (All.lookup i es) ≤ exprs es + lookup-≤ 0F (e ∷ es) = ∷ˡ-≤ e es + lookup-≤ (suc i) (e ∷ es) = ℕₚ.≤-trans (lookup-≤ i es) (∷ʳ-≤ e es) + + call>0 : ∀ (f : Function Σ Δ t) (es : All (Expression Σ Γ) Δ) → 0 < expr (call f es) + call>0 f es = ℕₚ.<-transˡ ℕₚ.0<1+n (ℕₚ.m≤n⊔m (exprs es) (suc (fun f))) + +module Cast where + expr : t ≡ t′ → Expression Σ Γ t → Expression Σ Γ t′ + expr refl e = e + + locRef : t ≡ t′ → LocalReference Σ Γ t → LocalReference Σ Γ t′ + locRef refl ref = ref + + homo-!! : ∀ (eq : t ≡ t′) (ref : LocalReference Σ Γ t) → expr eq (!! ref) ≡ !! (locRef eq ref) + homo-!! refl _ = refl + + expr-depth : ∀ (eq : t ≡ t′) (e : Expression Σ Γ t) → CallDepth.expr (expr eq e) ≡ CallDepth.expr e + expr-depth refl _ = refl + +module Elim where + expr : ∀ i → Expression Σ (insert Γ i t′) t → Expression Σ Γ t′ → Expression Σ Γ t + exprs : ∀ i → All (Expression Σ (insert Γ i t′)) ts → Expression Σ Γ t′ → All (Expression Σ Γ) ts + + expr i (lit x) e′ = lit x + expr i (state j) e′ = state j + expr i (var j) e′ with i Fin.≟ j + ... | yes refl = Cast.expr (sym (Vecₚ.insert-lookup _ i _)) e′ + ... | no i≢j = Cast.expr (punchOut-insert _ i≢j _) (var (punchOut i≢j)) + expr i (e ≟ e₁) e′ = expr i e e′ ≟ expr i e₁ e′ + expr i (e <? e₁) e′ = expr i e e′ <? expr i e₁ e′ + expr i (inv e) e′ = expr i e e′ + expr i (e && e₁) e′ = expr i e e′ && expr i e₁ e′ + expr i (e || e₁) e′ = expr i e e′ || expr i e₁ e′ + expr i (not e) e′ = not (expr i e e′) + expr i (e and e₁) e′ = expr i e e′ and expr i e₁ e′ + expr i (e or e₁) e′ = expr i e e′ or expr i e₁ e′ + expr i [ e ] e′ = [ expr i e e′ ] + expr i (unbox e) e′ = unbox (expr i e e′) + expr i (merge e e₁ e₂) e′ = merge (expr i e e′) (expr i e₁ e′) (expr i e₂ e′) + expr i (slice e e₁) e′ = slice (expr i e e′) (expr i e₁ e′) + expr i (cut e e₁) e′ = cut (expr i e e′) (expr i e₁ e′) + expr i (cast eq e) e′ = cast eq (expr i e e′) + expr i (- e) e′ = - expr i e e′ + expr i (e + e₁) e′ = expr i e e′ + expr i e₁ e′ + expr i (e * e₁) e′ = expr i e e′ * expr i e₁ e′ + expr i (e ^ x) e′ = expr i e e′ ^ x + expr i (e >> n) e′ = expr i e e′ >> n + expr i (rnd e) e′ = rnd (expr i e e′) + expr i (fin f e) e′ = fin f (expr i e e′) + expr i (asInt e) e′ = asInt (expr i e e′) + expr i nil e′ = nil + expr i (cons e e₁) e′ = cons (expr i e e′) (expr i e₁ e′) + expr i (head e) e′ = head (expr i e e′) + expr i (tail e) e′ = tail (expr i e e′) + expr i (call f es) e′ = call f (exprs i es e′) + expr i (if e then e₁ else e₂) e′ = if expr i e e′ then expr i e₁ e′ else expr i e₂ e′ + + exprs i [] e′ = [] + exprs i (e ∷ es) e′ = expr i e e′ ∷ exprs i es e′ + + expr-depth : ∀ i (e : Expression Σ _ t) (e′ : Expression Σ Γ t′) → CallDepth.expr (expr i e e′) ≤ CallDepth.expr e ⊔ CallDepth.expr e′ + exprs-depth : ∀ i (es : All (Expression Σ _) ts) (e′ : Expression Σ Γ t′) → CallDepth.exprs (exprs i es e′) ≤ CallDepth.exprs es ⊔ CallDepth.expr e′ + + expr-depth i (lit x) e′ = z≤n + expr-depth i (state j) e′ = z≤n + expr-depth i (var j) e′ with i Fin.≟ j + ... | yes refl = ℕₚ.≤-reflexive (Cast.expr-depth (sym (Vecₚ.insert-lookup _ i _)) e′) + ... | no i≢j = ℕₚ.≤-trans (ℕₚ.≤-reflexive (Cast.expr-depth (punchOut-insert _ i≢j _) (var (punchOut i≢j)))) z≤n + expr-depth i (e ≟ e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e <? e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (inv e) e′ = expr-depth i e e′ + expr-depth i (e && e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e || e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (not e) e′ = expr-depth i e e′ + expr-depth i (e and e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e or e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i [ e ] e′ = expr-depth i e e′ + expr-depth i (unbox e) e′ = expr-depth i e e′ + expr-depth i (merge e e₁ e₂) e′ = join-lubs (CallDepth.expr e′) 3 (expr-depth i e e′ , expr-depth i e₁ e′ , expr-depth i e₂ e′) + expr-depth i (slice e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (cut e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (cast eq e) e′ = expr-depth i e e′ + expr-depth i (- e) e′ = expr-depth i e e′ + expr-depth i (e + e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e * e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e ^ x) e′ = expr-depth i e e′ + expr-depth i (e >> n) e′ = expr-depth i e e′ + expr-depth i (rnd e) e′ = expr-depth i e e′ + expr-depth i (fin f e) e′ = expr-depth i e e′ + expr-depth i (asInt e) e′ = expr-depth i e e′ + expr-depth i nil e′ = z≤n + expr-depth i (cons e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (head e) e′ = expr-depth i e e′ + expr-depth i (tail e) e′ = expr-depth i e e′ + expr-depth i (call f es) e′ = join-lubs (CallDepth.expr e′) 2 (exprs-depth i es e′ , ℕₚ.m≤m⊔n _ (CallDepth.expr e′)) + expr-depth i (if e then e₁ else e₂) e′ = join-lubs (CallDepth.expr e′) 3 (expr-depth i e e′ , expr-depth i e₁ e′ , expr-depth i e₂ e′) + + exprs-depth i [] e′ = z≤n + exprs-depth i (e ∷ es) e′ = join-lubs (CallDepth.expr e′) 2 (exprs-depth i es e′ , expr-depth i e e′) + +module Weaken where + expr : ∀ i t′ → Expression Σ Γ t → Expression Σ (insert Γ i t′) t + exprs : ∀ i t′ → All (Expression Σ Γ) ts → All (Expression Σ (insert Γ i t′)) ts + + expr i t′ (lit x) = lit x + expr i t′ (state j) = state j + expr i t′ (var j) = Cast.expr (Vecₚ.insert-punchIn _ i t′ j) (var (punchIn i j)) + expr i t′ (e ≟ e₁) = expr i t′ e ≟ expr i t′ e₁ + expr i t′ (e <? e₁) = expr i t′ e <? expr i t′ e₁ + expr i t′ (inv e) = inv (expr i t′ e) + expr i t′ (e && e₁) = expr i t′ e && expr i t′ e₁ + expr i t′ (e || e₁) = expr i t′ e || expr i t′ e₁ + expr i t′ (not e) = not (expr i t′ e) + expr i t′ (e and e₁) = expr i t′ e and expr i t′ e₁ + expr i t′ (e or e₁) = expr i t′ e or expr i t′ e₁ + expr i t′ [ e ] = [ expr i t′ e ] + expr i t′ (unbox e) = unbox (expr i t′ e) + expr i t′ (merge e e₁ e₂) = merge (expr i t′ e) (expr i t′ e₁) (expr i t′ e₂) + expr i t′ (slice e e₁) = slice (expr i t′ e) (expr i t′ e₁) + expr i t′ (cut e e₁) = cut (expr i t′ e) (expr i t′ e₁) + expr i t′ (cast eq e) = cast eq (expr i t′ e) + expr i t′ (- e) = - expr i t′ e + expr i t′ (e + e₁) = expr i t′ e + expr i t′ e₁ + expr i t′ (e * e₁) = expr i t′ e * expr i t′ e₁ + expr i t′ (e ^ x) = expr i t′ e ^ x + expr i t′ (e >> n) = expr i t′ e >> n + expr i t′ (rnd e) = rnd (expr i t′ e) + expr i t′ (fin f e) = fin f (expr i t′ e) + expr i t′ (asInt e) = asInt (expr i t′ e) + expr i t′ nil = nil + expr i t′ (cons e e₁) = cons (expr i t′ e) (expr i t′ e₁) + expr i t′ (head e) = head (expr i t′ e) + expr i t′ (tail e) = tail (expr i t′ e) + expr i t′ (call f es) = call f (exprs i t′ es) + expr i t′ (if e then e₁ else e₂) = if expr i t′ e then expr i t′ e₁ else expr i t′ e₂ + + exprs i t′ [] = [] + exprs i t′ (e ∷ es) = expr i t′ e ∷ exprs i t′ es + + locRef : ∀ i t′ → LocalReference Σ Γ t → LocalReference Σ (insert Γ i t′) t + locRef i t′ (var j) = Cast.locRef (Vecₚ.insert-punchIn _ i t′ j) (var (punchIn i j)) + locRef i t′ [ ref ] = [ locRef i t′ ref ] + locRef i t′ (unbox ref) = unbox (locRef i t′ ref) + locRef i t′ (merge ref ref₁ e) = merge (locRef i t′ ref) (locRef i t′ ref₁) (expr i t′ e) + locRef i t′ (slice ref e) = slice (locRef i t′ ref) (expr i t′ e) + locRef i t′ (cut ref e) = cut (locRef i t′ ref) (expr i t′ e) + locRef i t′ (cast eq ref) = cast eq (locRef i t′ ref) + locRef i t′ nil = nil + locRef i t′ (cons ref ref₁) = cons (locRef i t′ ref) (locRef i t′ ref₁) + locRef i t′ (head ref) = head (locRef i t′ ref) + locRef i t′ (tail ref) = tail (locRef i t′ ref) + + locStmt : ∀ i t′ → LocalStatement Σ Γ → LocalStatement Σ (insert Γ i t′) + locStmt i t′ (s ∙ s₁) = locStmt i t′ s ∙ locStmt i t′ s₁ + locStmt i t′ skip = skip + locStmt i t′ (ref ≔ val) = locRef i t′ ref ≔ expr i t′ val + locStmt i t′ (declare e s) = declare (expr i t′ e) (locStmt (suc i) t′ s) + locStmt i t′ (if x then s) = if expr i t′ x then locStmt i t′ s + locStmt i t′ (if x then s else s₁) = if expr i t′ x then locStmt i t′ s else locStmt i t′ s₁ + locStmt i t′ (for n s) = for n (locStmt (suc i) t′ s) + + homo-!! : ∀ i t′ (ref : LocalReference Σ Γ t) → expr i t′ (!! ref) ≡ !! (locRef i t′ ref) + homo-!! i t′ (var j) = Cast.homo-!! (Vecₚ.insert-punchIn _ i t′ j) (var (punchIn i j)) + homo-!! i t′ [ ref ] = cong [_] (homo-!! i t′ ref) + homo-!! i t′ (unbox ref) = cong unbox (homo-!! i t′ ref) + homo-!! i t′ (merge ref ref₁ e) = cong₂ (λ x y → merge x y _) (homo-!! i t′ ref) (homo-!! i t′ ref₁) + homo-!! i t′ (slice ref e) = cong (λ x → slice x _) (homo-!! i t′ ref) + homo-!! i t′ (cut ref e) = cong (λ x → cut x _) (homo-!! i t′ ref) + homo-!! i t′ (cast eq ref) = cong (cast eq) (homo-!! i t′ ref) + homo-!! i t′ nil = refl + homo-!! i t′ (cons ref ref₁) = cong₂ cons (homo-!! i t′ ref) (homo-!! i t′ ref₁) + homo-!! i t′ (head ref) = cong head (homo-!! i t′ ref) + homo-!! i t′ (tail ref) = cong tail (homo-!! i t′ ref) + + expr-depth : ∀ i t′ (e : Expression Σ Γ t) → CallDepth.expr (expr i t′ e) ≡ CallDepth.expr e + exprs-depth : ∀ i t′ (es : All (Expression Σ Γ) ts) → CallDepth.exprs (exprs i t′ es) ≡ CallDepth.exprs es + + expr-depth i t′ (lit x) = refl + expr-depth i t′ (state j) = refl + expr-depth i t′ (var j) = Cast.expr-depth (Vecₚ.insert-punchIn _ i t′ j) (var (punchIn i j)) + expr-depth i t′ (e ≟ e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (e <? e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (inv e) = expr-depth i t′ e + expr-depth i t′ (e && e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (e || e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (not e) = expr-depth i t′ e + expr-depth i t′ (e and e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (e or e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ [ e ] = expr-depth i t′ e + expr-depth i t′ (unbox e) = expr-depth i t′ e + expr-depth i t′ (merge e e₁ e₂) = congₙ 3 (λ a b c → a ⊔ b ⊔ c) (expr-depth i t′ e) (expr-depth i t′ e₁) (expr-depth i t′ e₂) + expr-depth i t′ (slice e e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (cut e e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (cast eq e) = expr-depth i t′ e + expr-depth i t′ (- e) = expr-depth i t′ e + expr-depth i t′ (e + e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (e * e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (e ^ x) = expr-depth i t′ e + expr-depth i t′ (e >> n) = expr-depth i t′ e + expr-depth i t′ (rnd e) = expr-depth i t′ e + expr-depth i t′ (fin f e) = expr-depth i t′ e + expr-depth i t′ (asInt e) = expr-depth i t′ e + expr-depth i t′ nil = refl + expr-depth i t′ (cons e e₁) = cong₂ _⊔_ (expr-depth i t′ e) (expr-depth i t′ e₁) + expr-depth i t′ (head e) = expr-depth i t′ e + expr-depth i t′ (tail e) = expr-depth i t′ e + expr-depth i t′ (call f es) = cong (_⊔ _) (exprs-depth i t′ es) + expr-depth i t′ (if e then e₁ else e₂) = congₙ 3 (λ a b c → a ⊔ b ⊔ c) (expr-depth i t′ e) (expr-depth i t′ e₁) (expr-depth i t′ e₂) + + exprs-depth i t′ [] = refl + exprs-depth i t′ (e ∷ es) = cong₂ _⊔_ (exprs-depth i t′ es) (expr-depth i t′ e) + + locRef-depth : ∀ i t′ (ref : LocalReference Σ Γ t) → CallDepth.locRef (locRef i t′ ref) ≡ CallDepth.locRef ref + locRef-depth i t′ ref = begin-equality + CallDepth.locRef (locRef i t′ ref) ≡˘⟨ CallDepth.homo-!! (locRef i t′ ref) ⟩ + CallDepth.expr (!! (locRef i t′ ref)) ≡˘⟨ cong CallDepth.expr (homo-!! i t′ ref) ⟩ + CallDepth.expr (expr i t′ (!! ref)) ≡⟨ expr-depth i t′ (!! ref) ⟩ + CallDepth.expr (!! ref) ≡⟨ CallDepth.homo-!! ref ⟩ + CallDepth.locRef ref ∎ + + locStmt-depth : ∀ i t′ (s : LocalStatement Σ Γ) → CallDepth.locStmt (locStmt i t′ s) ≡ CallDepth.locStmt s + locStmt-depth i t′ (s ∙ s₁) = cong₂ _⊔_ (locStmt-depth i t′ s) (locStmt-depth i t′ s₁) + locStmt-depth i t′ skip = refl + locStmt-depth i t′ (ref ≔ val) = cong₂ _⊔_ (locRef-depth i t′ ref) (expr-depth i t′ val) + locStmt-depth i t′ (declare e s) = cong₂ _⊔_ (locStmt-depth (suc i) t′ s) (expr-depth i t′ e) + locStmt-depth i t′ (if x then s) = cong₂ _⊔_ (locStmt-depth i t′ s) (expr-depth i t′ x) + locStmt-depth i t′ (if x then s else s₁) = congₙ 3 (λ a b c → a ⊔ b ⊔ c) (locStmt-depth i t′ s) (locStmt-depth i t′ s₁) (expr-depth i t′ x) + locStmt-depth i t′ (for n s) = locStmt-depth (suc i) t′ s + +module Subst where + expr : ∀ i → Expression Σ Γ t → Expression Σ Γ (lookup Γ i) → Expression Σ Γ t + exprs : ∀ i → All (Expression Σ Γ) ts → Expression Σ Γ (lookup Γ i) → All (Expression Σ Γ) ts + + expr i (lit x) e′ = lit x + expr i (state j) e′ = state j + expr i (var j) e′ with i Fin.≟ j + ... | yes refl = e′ + ... | no i≢j = var j + expr i (e ≟ e₁) e′ = expr i e e′ ≟ expr i e₁ e′ + expr i (e <? e₁) e′ = expr i e e′ <? expr i e₁ e′ + expr i (inv e) e′ = inv (expr i e e′) + expr i (e && e₁) e′ = expr i e e′ && expr i e₁ e′ + expr i (e || e₁) e′ = expr i e e′ || expr i e₁ e′ + expr i (not e) e′ = not (expr i e e′) + expr i (e and e₁) e′ = expr i e e′ and expr i e₁ e′ + expr i (e or e₁) e′ = expr i e e′ or expr i e₁ e′ + expr i [ e ] e′ = [ expr i e e′ ] + expr i (unbox e) e′ = unbox (expr i e e′) + expr i (merge e e₁ e₂) e′ = merge (expr i e e′) (expr i e₁ e′) (expr i e₂ e′) + expr i (slice e e₁) e′ = slice (expr i e e′) (expr i e₁ e′) + expr i (cut e e₁) e′ = cut (expr i e e′) (expr i e₁ e′) + expr i (cast eq e) e′ = cast eq (expr i e e′) + expr i (- e) e′ = - expr i e e′ + expr i (e + e₁) e′ = expr i e e′ + expr i e₁ e′ + expr i (e * e₁) e′ = expr i e e′ * expr i e₁ e′ + expr i (e ^ x) e′ = expr i e e′ ^ x + expr i (e >> n) e′ = expr i e e′ >> n + expr i (rnd e) e′ = rnd (expr i e e′) + expr i (fin f e) e′ = fin f (expr i e e′) + expr i (asInt e) e′ = asInt (expr i e e′) + expr i nil e′ = nil + expr i (cons e e₁) e′ = cons (expr i e e′) (expr i e₁ e′) + expr i (head e) e′ = head (expr i e e′) + expr i (tail e) e′ = tail (expr i e e′) + expr i (call f es) e′ = call f (exprs i es e′) + expr i (if e then e₁ else e₂) e′ = if expr i e e′ then expr i e₁ e′ else expr i e₂ e′ + + exprs i [] e′ = [] + exprs i (e ∷ es) e′ = expr i e e′ ∷ exprs i es e′ + + expr-depth : ∀ i (e : Expression Σ Γ t) e′ → CallDepth.expr (expr i e e′) ≤ CallDepth.expr e ⊔ CallDepth.expr e′ + exprs-depth : ∀ i (es : All (Expression Σ Γ) ts) e′ → CallDepth.exprs (exprs i es e′) ≤ CallDepth.exprs es ⊔ CallDepth.expr e′ + + expr-depth i (lit x) e′ = z≤n + expr-depth i (state j) e′ = z≤n + expr-depth i (var j) e′ with i Fin.≟ j + ... | yes refl = ℕₚ.≤-refl + ... | no i≢j = z≤n + expr-depth i (e ≟ e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e <? e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (inv e) e′ = expr-depth i e e′ + expr-depth i (e && e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e || e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (not e) e′ = expr-depth i e e′ + expr-depth i (e and e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e or e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i [ e ] e′ = expr-depth i e e′ + expr-depth i (unbox e) e′ = expr-depth i e e′ + expr-depth i (merge e e₁ e₂) e′ = join-lubs (CallDepth.expr e′) 3 (expr-depth i e e′ , expr-depth i e₁ e′ , expr-depth i e₂ e′) + expr-depth i (slice e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (cut e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (cast eq e) e′ = expr-depth i e e′ + expr-depth i (- e) e′ = expr-depth i e e′ + expr-depth i (e + e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e * e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (e ^ x) e′ = expr-depth i e e′ + expr-depth i (e >> n) e′ = expr-depth i e e′ + expr-depth i (rnd e) e′ = expr-depth i e e′ + expr-depth i (fin f e) e′ = expr-depth i e e′ + expr-depth i (asInt e) e′ = expr-depth i e e′ + expr-depth i nil e′ = z≤n + expr-depth i (cons e e₁) e′ = join-lubs (CallDepth.expr e′) 2 (expr-depth i e e′ , expr-depth i e₁ e′) + expr-depth i (head e) e′ = expr-depth i e e′ + expr-depth i (tail e) e′ = expr-depth i e e′ + expr-depth i (call f es) e′ = join-lubs (CallDepth.expr e′) 2 (exprs-depth i es e′ , ℕₚ.m≤m⊔n _ (CallDepth.expr e′)) + expr-depth i (if e then e₁ else e₂) e′ = join-lubs (CallDepth.expr e′) 3 (expr-depth i e e′ , expr-depth i e₁ e′ , expr-depth i e₂ e′) + + exprs-depth i [] e′ = z≤n + exprs-depth i (e ∷ es) e′ = join-lubs (CallDepth.expr e′) 2 (exprs-depth i es e′ , expr-depth i e e′) + +module SubstAll where + expr : Expression Σ Γ t → All (Expression Σ Δ) Γ → Expression Σ Δ t + exprs : All (Expression Σ Γ) ts → All (Expression Σ Δ) Γ → All (Expression Σ Δ) ts + + expr (lit x) es′ = lit x + expr (state j) es′ = state j + expr (var j) es′ = All.lookup j es′ + expr (e ≟ e₁) es′ = expr e es′ ≟ expr e₁ es′ + expr (e <? e₁) es′ = expr e es′ <? expr e₁ es′ + expr (inv e) es′ = inv (expr e es′) + expr (e && e₁) es′ = expr e es′ && expr e₁ es′ + expr (e || e₁) es′ = expr e es′ || expr e₁ es′ + expr (not e) es′ = not (expr e es′) + expr (e and e₁) es′ = expr e es′ and expr e₁ es′ + expr (e or e₁) es′ = expr e es′ or expr e₁ es′ + expr [ e ] es′ = [ expr e es′ ] + expr (unbox e) es′ = unbox (expr e es′) + expr (merge e e₁ e₂) es′ = merge (expr e es′) (expr e₁ es′) (expr e₂ es′) + expr (slice e e₁) es′ = slice (expr e es′) (expr e₁ es′) + expr (cut e e₁) es′ = cut (expr e es′) (expr e₁ es′) + expr (cast eq e) es′ = cast eq (expr e es′) + expr (- e) es′ = - expr e es′ + expr (e + e₁) es′ = expr e es′ + expr e₁ es′ + expr (e * e₁) es′ = expr e es′ * expr e₁ es′ + expr (e ^ x) es′ = expr e es′ ^ x + expr (e >> n) es′ = expr e es′ >> n + expr (rnd e) es′ = rnd (expr e es′) + expr (fin f e) es′ = fin f (expr e es′) + expr (asInt e) es′ = asInt (expr e es′) + expr nil es′ = nil + expr (cons e e₁) es′ = cons (expr e es′) (expr e₁ es′) + expr (head e) es′ = head (expr e es′) + expr (tail e) es′ = tail (expr e es′) + expr (call f es) es′ = call f (exprs es es′) + expr (if e then e₁ else e₂) es′ = if expr e es′ then expr e₁ es′ else expr e₂ es′ + + exprs [] es′ = [] + exprs (e ∷ es) es′ = expr e es′ ∷ exprs es es′ + + expr-depth : ∀ (e : Expression Σ Γ t) (es′ : All (Expression Σ Δ) Γ) → CallDepth.expr (expr e es′) ≤ CallDepth.expr e ⊔ CallDepth.exprs es′ + exprs-depth : ∀ (es : All (Expression Σ Γ) ts) (es′ : All (Expression Σ Δ) Γ) → CallDepth.exprs (exprs es es′) ≤ CallDepth.exprs es ⊔ CallDepth.exprs es′ + + expr-depth (lit x) es′ = z≤n + expr-depth (state j) es′ = z≤n + expr-depth (var j) es′ = CallDepth.lookup-≤ j es′ + expr-depth (e ≟ e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (e <? e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (inv e) es′ = expr-depth e es′ + expr-depth (e && e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (e || e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (not e) es′ = expr-depth e es′ + expr-depth (e and e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (e or e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth [ e ] es′ = expr-depth e es′ + expr-depth (unbox e) es′ = expr-depth e es′ + expr-depth (merge e e₁ e₂) es′ = join-lubs (CallDepth.exprs es′) 3 (expr-depth e es′ , expr-depth e₁ es′ , expr-depth e₂ es′) + expr-depth (slice e e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (cut e e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (cast eq e) es′ = expr-depth e es′ + expr-depth (- e) es′ = expr-depth e es′ + expr-depth (e + e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (e * e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (e ^ x) es′ = expr-depth e es′ + expr-depth (e >> n) es′ = expr-depth e es′ + expr-depth (rnd e) es′ = expr-depth e es′ + expr-depth (fin f e) es′ = expr-depth e es′ + expr-depth (asInt e) es′ = expr-depth e es′ + expr-depth nil es′ = z≤n + expr-depth (cons e e₁) es′ = join-lubs (CallDepth.exprs es′) 2 (expr-depth e es′ , expr-depth e₁ es′) + expr-depth (head e) es′ = expr-depth e es′ + expr-depth (tail e) es′ = expr-depth e es′ + expr-depth (call f es) es′ = join-lubs (CallDepth.exprs es′) 2 (exprs-depth es es′ , ℕₚ.m≤m⊔n _ (CallDepth.exprs es′)) + expr-depth (if e then e₁ else e₂) es′ = join-lubs (CallDepth.exprs es′) 3 (expr-depth e es′ , expr-depth e₁ es′ , expr-depth e₂ es′) + + exprs-depth [] es′ = z≤n + exprs-depth (e ∷ es) es′ = join-lubs (CallDepth.exprs es′) 2 (exprs-depth es es′ , expr-depth e es′) + +module Update where + expr : LocalReference Σ Γ t → Expression Σ Γ t → Expression Σ Γ t′ → Expression Σ Γ t′ + expr (var i) val e′ = Subst.expr i e′ val + expr [ ref ] val e′ = expr ref (unbox val) e′ + expr (unbox ref) val e′ = expr ref [ val ] e′ + expr (merge ref ref₁ e) val e′ = expr ref₁ (cut val e) (expr ref (slice val e) e′) + expr (slice ref e) val e′ = expr ref (merge val (cut (!! ref) e) e) e′ + expr (cut ref e) val e′ = expr ref (merge (slice (!! ref) e) val e) e′ + expr (cast eq ref) val e′ = expr ref (cast (sym eq) val) e′ + expr nil val e′ = e′ + expr (cons ref ref₁) val e′ = expr ref₁ (tail val) (expr ref (head val) e′) + expr (head ref) val e′ = expr ref (cons val (tail (!! ref))) e′ + expr (tail ref) val e′ = expr ref (cons (head (!! ref)) val) e′ + + expr-depth : ∀ (ref : LocalReference Σ Γ t) val (e′ : Expression Σ Γ t′) → CallDepth.expr (expr ref val e′) ≤ CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val) + expr-depth (var i) val e′ = Subst.expr-depth i e′ val + expr-depth [ ref ] val e′ = expr-depth ref (unbox val) e′ + expr-depth (unbox ref) val e′ = expr-depth ref [ val ] e′ + expr-depth (merge ref ref₁ e) val e′ = begin + CallDepth.expr (expr ref₁ (cut val e) (expr ref (slice val e) e′)) + ≤⟨ expr-depth ref₁ _ _ ⟩ + CallDepth.expr (expr ref (slice val e) e′) ⊔ (CallDepth.locRef ref₁ ⊔ (CallDepth.expr val ⊔ CallDepth.expr e)) + ≤⟨ ℕₚ.⊔-monoˡ-≤ _ (expr-depth ref (slice val e) e′) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ CallDepth.expr e)) ⊔ (CallDepth.locRef ref₁ ⊔ (CallDepth.expr val ⊔ CallDepth.expr e)) + ≡⟨ solve-⊔ 5 (λ a b c d e → (a ⊕ (b ⊕ (e ⊕ d))) ⊕ (c ⊕ (e ⊕ d)) ⊜ a ⊕ (((b ⊕ c) ⊕ d) ⊕ e)) refl (CallDepth.expr e′) (CallDepth.locRef ref) _ _ _ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.locRef ref₁ ⊔ CallDepth.expr e ⊔ CallDepth.expr val) ∎ -elimAt-pres-callDepth i (abort e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (e ≟ e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (e <? e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (inv e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (e && e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (e || e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (not e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (e and e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (e or e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i [ e ] e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (unbox e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (splice e e₁ e₂) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) ⊔ callDepth (elimAt i e₂ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′)) (elimAt-pres-callDepth i e₂ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) ⊔ (callDepth e′ ⊔ callDepth e₂) - ≡⟨ ⊔-solve 4 (λ m n o p → (((m ⊕ n) ⊕ (m ⊕ o)) ⊕ (m ⊕ p)) ⊜ (m ⊕ ((n ⊕ o) ⊕ p))) refl (callDepth e′) (callDepth e) (callDepth e₁) (callDepth e₂) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂) - ∎ -elimAt-pres-callDepth i (cut e e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (cast eq e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (- e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (e + e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (e * e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (e ^ x) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (e >> x) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (rnd e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (fin x e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (asInt e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i nil e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -elimAt-pres-callDepth i (cons e e₁) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ m n o → (m ⊕ n) ⊕ (m ⊕ o) ⊜ m ⊕ (n ⊕ o)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -elimAt-pres-callDepth i (head e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (tail e) e′ = elimAt-pres-callDepth i e e′ -elimAt-pres-callDepth i (call f es) e′ = begin - suc (funCallDepth f) ⊔ callDepth′ (elimAt′ i es e′) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (suc (funCallDepth f)) (elimAt′-pres-callDepth i es e′) ⟩ - suc (funCallDepth f) ⊔ (callDepth e′ ⊔ callDepth′ es) - ≡⟨ ⊔-solve 3 (λ a b c → b ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (suc (funCallDepth f)) (callDepth′ es) ⟩ - callDepth e′ ⊔ (suc (funCallDepth f) ⊔ callDepth′ es) - ∎ -elimAt-pres-callDepth i (if e then e₁ else e₂) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth (elimAt i e₁ e′) ⊔ callDepth (elimAt i e₂ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt-pres-callDepth i e₁ e′)) (elimAt-pres-callDepth i e₂ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) ⊔ (callDepth e′ ⊔ callDepth e₂) - ≡⟨ ⊔-solve 4 (λ m n o p → (((m ⊕ n) ⊕ (m ⊕ o)) ⊕ (m ⊕ p)) ⊜ (m ⊕ ((n ⊕ o) ⊕ p))) refl (callDepth e′) (callDepth e) (callDepth e₁) (callDepth e₂) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂) - ∎ - -elimAt′-pres-callDepth i [] e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -elimAt′-pres-callDepth i (e ∷ es) e′ = begin - callDepth (elimAt i e e′) ⊔ callDepth′ (elimAt′ i es e′) - ≤⟨ ℕₚ.⊔-mono-≤ (elimAt-pres-callDepth i e e′) (elimAt′-pres-callDepth i es e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth′ es) - ≡⟨ ⊔-solve 3 (λ a b c → ((a ⊕ b) ⊕ (a ⊕ c)) ⊜ (a ⊕ (b ⊕ c))) refl (callDepth e′) (callDepth e) (callDepth′ es) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth′ es) - ∎ - -wknAt-pres-callDepth : ∀ i (e : Expression Γ t) → callDepth (wknAt {t′ = t′} i e) ≡ callDepth e -wknAt′-pres-callDepth : ∀ i (es : All (Expression Γ) Δ) → callDepth′ (wknAt′ {t′ = t′} i es) ≡ callDepth′ es - -wknAt-pres-callDepth i (Code.lit x) = refl -wknAt-pres-callDepth i (Code.state j) = refl -wknAt-pres-callDepth {Γ = Γ} i (Code.var j) = castType-pres-callDepth (var {Γ = Vec.insert Γ i _} (Fin.punchIn i j)) (Vecₚ.insert-punchIn Γ i _ j) -wknAt-pres-callDepth i (Code.abort e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (e Code.≟ e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (e Code.<? e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (Code.inv e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (e Code.&& e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (e Code.|| e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (Code.not e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (e Code.and e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (e Code.or e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i Code.[ e ] = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.unbox e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.splice e e₁ e₂) = congₙ 3 (λ m n o → m ⊔ n ⊔ o) (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) (wknAt-pres-callDepth i e₂) -wknAt-pres-callDepth i (Code.cut e e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (Code.cast eq e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.- e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (e Code.+ e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (e Code.* e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (e Code.^ x) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (e Code.>> x) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.rnd e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.fin x e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.asInt e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i Code.nil = refl -wknAt-pres-callDepth i (Code.cons e e₁) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) -wknAt-pres-callDepth i (Code.head e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.tail e) = wknAt-pres-callDepth i e -wknAt-pres-callDepth i (Code.call f es) = cong (suc (funCallDepth f) ⊔_) (wknAt′-pres-callDepth i es) -wknAt-pres-callDepth i (Code.if e then e₁ else e₂) = congₙ 3 (λ m n o → m ⊔ n ⊔ o) (wknAt-pres-callDepth i e) (wknAt-pres-callDepth i e₁) (wknAt-pres-callDepth i e₂) - -wknAt′-pres-callDepth i [] = refl -wknAt′-pres-callDepth i (e ∷ es) = cong₂ _⊔_ (wknAt-pres-callDepth i e) (wknAt′-pres-callDepth i es) - -substAt-pres-callDepth : ∀ i (e : Expression Γ t) e′ → callDepth (substAt i e e′) ℕ.≤ callDepth e′ ⊔ callDepth e -substAt-pres-callDepth i (Code.lit x) e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -substAt-pres-callDepth i (Code.state j) e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -substAt-pres-callDepth i (Code.var j) e′ with i Fin.≟ j -... | yes refl = ℕₚ.m≤m⊔n (callDepth e′) 0 -... | no _ = ℕₚ.m≤n⊔m (callDepth e′) 0 -substAt-pres-callDepth i (Code.abort e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (e Code.≟ e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (e Code.<? e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (Code.inv e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (e Code.&& e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (e Code.|| e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (Code.not e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (e Code.and e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (e Code.or e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i Code.[ e ] e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.unbox e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.splice e e₁ e₂) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) ⊔ callDepth (substAt i e₂ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′)) (substAt-pres-callDepth i e₂ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) ⊔ (callDepth e′ ⊔ callDepth e₂) - ≡⟨ ⊔-solve 4 (λ a b c d → ((a ⊕ b) ⊕ (a ⊕ c)) ⊕ (a ⊕ d) ⊜ a ⊕ ((b ⊕ c) ⊕ d)) refl (callDepth e′) (callDepth e) (callDepth e₁) (callDepth e₂) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂) - ∎ -substAt-pres-callDepth i (Code.cut e e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (Code.cast eq e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.- e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (e Code.+ e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (e Code.* e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (e Code.^ x) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (e Code.>> x) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.rnd e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.fin x e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.asInt e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i Code.nil e′ = ℕₚ.m≤n⊔m (callDepth e′) 0 -substAt-pres-callDepth i (Code.cons e e₁) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ b) ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (callDepth e) (callDepth e₁) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁) - ∎ -substAt-pres-callDepth i (Code.head e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.tail e) e′ = substAt-pres-callDepth i e e′ -substAt-pres-callDepth i (Code.call f es) e′ = begin - suc (funCallDepth f) ⊔ callDepth′ (substAt′ i es e′) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (suc (funCallDepth f)) (helper es) ⟩ - suc (funCallDepth f) ⊔ (callDepth e′ ⊔ callDepth′ es) - ≡⟨ ⊔-solve 3 (λ a b c → b ⊕ (a ⊕ c) ⊜ a ⊕ (b ⊕ c)) refl (callDepth e′) (suc (funCallDepth f)) (callDepth′ es) ⟩ - callDepth e′ ⊔ (suc (funCallDepth f) ⊔ callDepth′ es) - ∎ - where - helper : ∀ {n ts} (es : All _ {n} ts) → callDepth′ (substAt′ i es e′) ℕ.≤ callDepth e′ ⊔ callDepth′ es - helper [] = ℕₚ.m≤n⊔m (callDepth e′) 0 - helper (e ∷ es) = begin - callDepth (substAt i e e′) ⊔ callDepth′ (substAt′ i es e′) - ≤⟨ ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (helper es) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth′ es) - ≡⟨ ⊔-solve 3 (λ a b c → ((a ⊕ b) ⊕ (a ⊕ c)) ⊜ (a ⊕ (b ⊕ c))) refl (callDepth e′) (callDepth e) (callDepth′ es) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth′ es) + expr-depth (slice ref e) val e′ = begin + CallDepth.expr (expr ref (merge val (cut (!! ref) e) e) e′) + ≤⟨ expr-depth ref (merge val (cut (!! ref) e) e) e′ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ (CallDepth.expr (!! ref) ⊔ CallDepth.expr e) ⊔ CallDepth.expr e)) + ≡⟨ cong (λ x → CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ (x ⊔ _) ⊔ _))) (CallDepth.homo-!! ref) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr e) ⊔ CallDepth.expr e)) + ≡⟨ cong (CallDepth.expr e′ ⊔_) $ solve-⊔ 3 (λ a b c → a ⊕ ((c ⊕ (a ⊕ b)) ⊕ b) ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locRef ref) _ _ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr e ⊔ CallDepth.expr val) ∎ -substAt-pres-callDepth i (Code.if e then e₁ else e₂) e′ = begin - callDepth (substAt i e e′) ⊔ callDepth (substAt i e₁ e′) ⊔ callDepth (substAt i e₂ e′) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (substAt-pres-callDepth i e e′) (substAt-pres-callDepth i e₁ e′)) (substAt-pres-callDepth i e₂ e′) ⟩ - callDepth e′ ⊔ callDepth e ⊔ (callDepth e′ ⊔ callDepth e₁) ⊔ (callDepth e′ ⊔ callDepth e₂) - ≡⟨ ⊔-solve 4 (λ a b c d → ((a ⊕ b) ⊕ (a ⊕ c)) ⊕ (a ⊕ d) ⊜ a ⊕ ((b ⊕ c) ⊕ d)) refl (callDepth e′) (callDepth e) (callDepth e₁) (callDepth e₂) ⟩ - callDepth e′ ⊔ (callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂) - ∎ - -updateRef-pres-callDepth : ∀ {e : Expression Γ t} ref stateless val (e′ : Expression Γ t′) → - callDepth (updateRef {e = e} ref stateless val e′) ℕ.≤ callDepth e ⊔ callDepth val ⊔ callDepth e′ -updateRef-pres-callDepth (state i) stateless val e′ = contradiction (state i) stateless -updateRef-pres-callDepth (var i) stateless val e′ = substAt-pres-callDepth i e′ val -updateRef-pres-callDepth (abort e) stateless val e′ = ℕₚ.m≤n⊔m (callDepth e ⊔ callDepth val) (callDepth e′) -updateRef-pres-callDepth [ ref ] stateless val e′ = updateRef-pres-callDepth ref (stateless ∘ [_]) (unbox val) e′ -updateRef-pres-callDepth (unbox ref) stateless val e′ = updateRef-pres-callDepth ref (stateless ∘ unbox) [ val ] e′ -updateRef-pres-callDepth (splice {e₁ = e₁} {e₂ = e₂} ref ref₁ e₃) stateless val e′ = begin - callDepth outer - ≤⟨ updateRef-pres-callDepth ref₁ (stateless ∘ (λ x → spliceʳ _ x e₃)) (head (tail (cut val e₃))) inner ⟩ - callDepth e₂ ⊔ (callDepth val ⊔ callDepth e₃) ⊔ callDepth inner - ≤⟨ ℕₚ.⊔-monoʳ-≤ (callDepth e₂ ⊔ (callDepth val ⊔ callDepth e₃)) (updateRef-pres-callDepth ref (stateless ∘ (λ x → spliceˡ x _ e₃)) (head (cut val e₃)) e′) ⟩ - callDepth e₂ ⊔ (callDepth val ⊔ callDepth e₃) ⊔ (callDepth e₁ ⊔ (callDepth val ⊔ callDepth e₃) ⊔ callDepth e′) - ≡⟨ ⊔-solve 5 (λ a b c d e → ((b ⊕ (d ⊕ c)) ⊕ ((a ⊕ (d ⊕ c)) ⊕ e)) ⊜ ((((a ⊕ b) ⊕ c) ⊕ d) ⊕ e)) refl (callDepth e₁) (callDepth e₂) (callDepth e₃) (callDepth val) (callDepth e′) ⟩ - callDepth e₁ ⊔ callDepth e₂ ⊔ callDepth e₃ ⊔ callDepth val ⊔ callDepth e′ - ∎ - where - inner = updateRef ref (stateless ∘ (λ x → spliceˡ x _ e₃)) (head (cut val e₃)) e′ - outer = updateRef ref₁ (stateless ∘ (λ x → spliceʳ _ x e₃)) (head (tail (cut val e₃))) inner -updateRef-pres-callDepth (cut {e₁ = e₁} ref e₂) stateless val e′ = begin - callDepth (updateRef ref (stateless ∘ (λ x → (cut x e₂))) (splice (head val) (head (tail val)) e₂) e′) - ≤⟨ updateRef-pres-callDepth ref (stateless ∘ (λ x → (cut x e₂))) (splice (head val) (head (tail val)) e₂) e′ ⟩ - callDepth e₁ ⊔ (callDepth val ⊔ callDepth val ⊔ callDepth e₂) ⊔ callDepth e′ - ≡⟨ ⊔-solve 4 (λ a b c d → (a ⊕ ((c ⊕ c) ⊕ b)) ⊕ d ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (callDepth e₁) (callDepth e₂) (callDepth val) (callDepth e′) ⟩ - callDepth e₁ ⊔ callDepth e₂ ⊔ callDepth val ⊔ callDepth e′ - ∎ -updateRef-pres-callDepth (cast eq ref) stateless val e′ = updateRef-pres-callDepth ref (stateless ∘ cast eq) (cast (sym eq) val) e′ -updateRef-pres-callDepth nil stateless val e′ = ℕₚ.m≤n⊔m (callDepth val) (callDepth e′) -updateRef-pres-callDepth (cons {e₁ = e₁} {e₂ = e₂} ref ref₁) stateless val e′ = begin - callDepth outer - ≤⟨ updateRef-pres-callDepth ref₁ (stateless ∘ consʳ e₁) (tail val) inner ⟩ - callDepth e₂ ⊔ callDepth val ⊔ callDepth inner - ≤⟨ ℕₚ.⊔-monoʳ-≤ (callDepth e₂ ⊔ callDepth val) (updateRef-pres-callDepth ref (stateless ∘ (λ x → consˡ x e₂)) (head val) e′) ⟩ - callDepth e₂ ⊔ callDepth val ⊔ (callDepth e₁ ⊔ callDepth val ⊔ callDepth e′) - ≡⟨ ⊔-solve 4 (λ a b c d → (b ⊕ c) ⊕ ((a ⊕ c) ⊕ d) ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (callDepth e₁) (callDepth e₂) (callDepth val) (callDepth e′) ⟩ - callDepth e₁ ⊔ callDepth e₂ ⊔ callDepth val ⊔ callDepth e′ - ∎ - where - inner = updateRef ref (stateless ∘ (λ x → consˡ x e₂)) (head val) e′ - outer = updateRef ref₁ (stateless ∘ consʳ e₁) (tail val) inner -updateRef-pres-callDepth (head {e = e} ref) stateless val e′ = begin - callDepth (updateRef ref (stateless ∘ head) (cons val (tail e)) e′) - ≤⟨ updateRef-pres-callDepth ref (stateless ∘ head) (cons val (tail e)) e′ ⟩ - callDepth e ⊔ (callDepth val ⊔ callDepth e) ⊔ callDepth e′ - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ (b ⊕ a)) ⊕ c ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth val) (callDepth e′) ⟩ - callDepth e ⊔ callDepth val ⊔ callDepth e′ - ∎ -updateRef-pres-callDepth (tail {e = e} ref) stateless val e′ = begin - callDepth (updateRef ref (stateless ∘ tail) (cons (head e) val) e′) - ≤⟨ updateRef-pres-callDepth ref (stateless ∘ tail) (cons (head e) val) e′ ⟩ - callDepth e ⊔ (callDepth e ⊔ callDepth val) ⊔ callDepth e′ - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ (a ⊕ b)) ⊕ c ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth val) (callDepth e′) ⟩ - callDepth e ⊔ callDepth val ⊔ callDepth e′ - ∎ - -subst-pres-callDepth : ∀ (e : Expression Γ t) (args : All (Expression Δ) Γ) → callDepth (subst e args) ℕ.≤ callDepth e ⊔ callDepth′ args -subst-pres-callDepth (lit x) args = ℕ.z≤n -subst-pres-callDepth (state i) args = ℕ.z≤n -subst-pres-callDepth (var i) args = helper i args - where - helper : ∀ i (es : All (Expression Γ) Δ) → callDepth (All.lookup i es) ℕ.≤ callDepth′ es - helper 0F (e ∷ es) = ℕₚ.m≤m⊔n (callDepth e) (callDepth′ es) - helper (suc i) (e ∷ es) = ℕₚ.m≤n⇒m≤o⊔n (callDepth e) (helper i es) -subst-pres-callDepth (abort e) args = subst-pres-callDepth e args -subst-pres-callDepth (e ≟ e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (e <? e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (inv e) args = subst-pres-callDepth e args -subst-pres-callDepth (e && e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (e || e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (not e) args = subst-pres-callDepth e args -subst-pres-callDepth (e and e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (e or e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth [ e ] args = subst-pres-callDepth e args -subst-pres-callDepth (unbox e) args = subst-pres-callDepth e args -subst-pres-callDepth (splice e e₁ e₂) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) ⊔ callDepth (subst e₂ args) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args)) (subst-pres-callDepth e₂ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) ⊔ (callDepth e₂ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 4 (λ a b c d → ((a ⊕ d) ⊕ (b ⊕ d)) ⊕ (c ⊕ d) ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (callDepth e) (callDepth e₁) (callDepth e₂) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (cut e e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (cast eq e) args = subst-pres-callDepth e args -subst-pres-callDepth (- e) args = subst-pres-callDepth e args -subst-pres-callDepth (e + e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (e * e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (e ^ x) args = subst-pres-callDepth e args -subst-pres-callDepth (e >> x) args = subst-pres-callDepth e args -subst-pres-callDepth (rnd e) args = subst-pres-callDepth e args -subst-pres-callDepth (fin x e) args = subst-pres-callDepth e args -subst-pres-callDepth (asInt e) args = subst-pres-callDepth e args -subst-pres-callDepth nil args = ℕ.z≤n -subst-pres-callDepth (cons e e₁) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth e₁) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth′ args - ∎ -subst-pres-callDepth (head e) args = subst-pres-callDepth e args -subst-pres-callDepth (tail e) args = subst-pres-callDepth e args -subst-pres-callDepth (call f es) args = begin - suc (funCallDepth f) ⊔ callDepth′ (subst′ es args) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (suc (funCallDepth f)) (helper es args) ⟩ - suc (funCallDepth f) ⊔ (callDepth′ es ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → a ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (suc (funCallDepth f)) (callDepth′ es) (callDepth′ args) ⟩ - suc (funCallDepth f) ⊔ callDepth′ es ⊔ callDepth′ args - ∎ - where - helper : ∀ {n ts} (es : All (Expression Γ) {n} ts) (args : All (Expression Δ) Γ) → callDepth′ (subst′ es args) ℕ.≤ callDepth′ es ℕ.⊔ callDepth′ args - helper [] args = ℕ.z≤n - helper (e ∷ es) args = begin - callDepth (subst e args) ⊔ callDepth′ (subst′ es args) - ≤⟨ ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (helper es args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth′ es ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → (a ⊕ c) ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth e) (callDepth′ es) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth′ es ⊔ callDepth′ args + expr-depth (cut ref e) val e′ = begin + CallDepth.expr (expr ref (merge (slice (!! ref) e) val e) e′) + ≤⟨ expr-depth ref (merge (slice (!! ref) e) val e) e′ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr (!! ref) ⊔ CallDepth.expr e ⊔ CallDepth.expr val ⊔ CallDepth.expr e)) + ≡⟨ cong (λ x → CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (x ⊔ _ ⊔ _ ⊔ _))) (CallDepth.homo-!! ref) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr e ⊔ CallDepth.expr val ⊔ CallDepth.expr e)) + ≡⟨ cong (CallDepth.expr e′ ⊔_) $ solve-⊔ 3 (λ a b c → a ⊕ (((a ⊕ b) ⊕ c) ⊕ b) ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locRef ref) _ _ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr e ⊔ CallDepth.expr val) + ∎ + expr-depth (cast eq ref) val e′ = expr-depth ref (cast (sym eq) val) e′ + expr-depth nil val e′ = ℕₚ.m≤m⊔n (CallDepth.expr e′) _ + expr-depth (cons ref ref₁) val e′ = begin + CallDepth.expr (expr ref₁ (tail val) (expr ref (head val) e′)) + ≤⟨ expr-depth ref₁ (tail val) (expr ref (head val) e′) ⟩ + CallDepth.expr (expr ref (head val) e′) ⊔ (CallDepth.locRef ref₁ ⊔ CallDepth.expr val) + ≤⟨ ℕₚ.⊔-monoˡ-≤ _ (expr-depth ref (head val) e′) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val) ⊔ (CallDepth.locRef ref₁ ⊔ CallDepth.expr val) + ≡⟨ solve-⊔ 4 (λ a b c d → (a ⊕ (b ⊕ d)) ⊕ (c ⊕ d) ⊜ a ⊕ ((b ⊕ c) ⊕ d)) refl (CallDepth.expr e′) (CallDepth.locRef ref) _ _ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.locRef ref₁ ⊔ CallDepth.expr val) + ∎ + expr-depth (head ref) val e′ = begin + CallDepth.expr (expr ref (cons val (tail (!! ref))) e′) + ≤⟨ expr-depth ref (cons val (tail (!! ref))) e′ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ CallDepth.expr (!! ref))) + ≡⟨ cong (λ x → CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ x))) (CallDepth.homo-!! ref) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr val ⊔ CallDepth.locRef ref)) + ≡⟨ cong (CallDepth.expr e′ ⊔_) (solve-⊔ 2 (λ a b → a ⊕ (b ⊕ a) ⊜ a ⊕ b) refl (CallDepth.locRef ref) _) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val) + ∎ + expr-depth (tail ref) val e′ = begin + CallDepth.expr (expr ref (cons (head (!! ref)) val) e′) + ≤⟨ expr-depth ref (cons (head (!! ref)) val) e′ ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.expr (!! ref) ⊔ CallDepth.expr val)) + ≡⟨ cong (λ x → CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (x ⊔ _))) (CallDepth.homo-!! ref) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val)) + ≡⟨ cong (CallDepth.expr e′ ⊔_) (solve-⊔ 2 (λ a b → a ⊕ (a ⊕ b) ⊜ a ⊕ b) refl (CallDepth.locRef ref) _) ⟩ + CallDepth.expr e′ ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val) ∎ -subst-pres-callDepth (if e then e₁ else e₂) args = begin - callDepth (subst e args) ⊔ callDepth (subst e₁ args) ⊔ callDepth (subst e₂ args) - ≤⟨ ℕₚ.⊔-mono-≤ (ℕₚ.⊔-mono-≤ (subst-pres-callDepth e args) (subst-pres-callDepth e₁ args)) (subst-pres-callDepth e₂ args) ⟩ - callDepth e ⊔ callDepth′ args ⊔ (callDepth e₁ ⊔ callDepth′ args) ⊔ (callDepth e₂ ⊔ callDepth′ args) - ≡⟨ ⊔-solve 4 (λ a b c d → ((a ⊕ d) ⊕ (b ⊕ d)) ⊕ (c ⊕ d) ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (callDepth e) (callDepth e₁) (callDepth e₂) (callDepth′ args) ⟩ - callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂ ⊔ callDepth′ args - ∎ - -wkn-pres-callDepth : ∀ t i (s : Statement Γ) → stmtCallDepth (wknStatementAt t i s) ≡ stmtCallDepth s -wkn-pres-callDepth t i (s Code.∙ s₁) = cong₂ _⊔_ (wkn-pres-callDepth t i s) (wkn-pres-callDepth t i s₁) -wkn-pres-callDepth t i Code.skip = refl -wkn-pres-callDepth t i (ref Code.≔ x) = cong₂ _⊔_ (wknAt-pres-callDepth i ref) (wknAt-pres-callDepth i x) -wkn-pres-callDepth t i (Code.declare x s) = cong₂ _⊔_ (wknAt-pres-callDepth i x) (wkn-pres-callDepth t (suc i) s) -wkn-pres-callDepth t i (Code.invoke p es) = cong (procCallDepth p ⊔_) (wknAt′-pres-callDepth i es) -wkn-pres-callDepth t i (Code.if x then s) = cong₂ _⊔_ (wknAt-pres-callDepth i x) (wkn-pres-callDepth t i s) -wkn-pres-callDepth t i (Code.if x then s else s₁) = congₙ 3 (λ m n o → m ⊔ n ⊔ o) (wknAt-pres-callDepth i x) (wkn-pres-callDepth t i s) (wkn-pres-callDepth t i s₁) -wkn-pres-callDepth t i (Code.for m s) = wkn-pres-callDepth t (suc i) s - -private - index₀ : Statement Γ → ℕ - index₀ (s ∙ s₁) = index₀ s ℕ.+ index₀ s₁ - index₀ skip = 0 - index₀ (ref ≔ x) = 0 - index₀ (declare x s) = index₀ s - index₀ (invoke p es) = 0 - index₀ (if x then s) = index₀ s - index₀ (if x then s else s₁) = suc (index₀ s ℕ.+ index₀ s₁) - index₀ (for m s) = index₀ s - - index₁ : Statement Γ → ℕ - index₁ (s ∙ s₁) = suc (index₁ s ℕ.+ index₁ s₁) - index₁ skip = 0 - index₁ (ref ≔ x) = 0 - index₁ (declare x s) = index₁ s - index₁ (invoke x x₁) = 0 - index₁ (if x then s) = suc (3 ℕ.* index₁ s) - index₁ (if x then s else s₁) = suc (3 ℕ.* index₁ s ℕ.+ index₁ s₁) - index₁ (for m s) = suc (index₁ s) - - index₂ : Statement Γ → ℕ - index₂ (s ∙ s₁) = 0 - index₂ skip = 0 - index₂ (ref ≔ x) = 0 - index₂ (declare x s) = suc (index₂ s) - index₂ (invoke _ _) = 0 - index₂ (if x then s) = 2 ℕ.* index₂ s - index₂ (if x then s else s₁) = 0 - index₂ (for m s) = 0 - - wkn-pres-index₀ : ∀ t i s → index₀ (wknStatementAt {Γ = Γ} t i s) ≡ index₀ s - wkn-pres-index₀ _ i (s ∙ s₁) = cong₂ ℕ._+_ (wkn-pres-index₀ _ i s) (wkn-pres-index₀ _ i s₁) - wkn-pres-index₀ _ i skip = refl - wkn-pres-index₀ _ i (ref ≔ x) = refl - wkn-pres-index₀ _ i (declare x s) = wkn-pres-index₀ _ (suc i) s - wkn-pres-index₀ _ i (invoke x x₁) = refl - wkn-pres-index₀ _ i (if x then s) = wkn-pres-index₀ _ i s - wkn-pres-index₀ _ i (if x then s else s₁) = cong₂ (λ m n → suc (m ℕ.+ n)) (wkn-pres-index₀ _ i s) (wkn-pres-index₀ _ i s₁) - wkn-pres-index₀ _ i (for m s) = wkn-pres-index₀ _ (suc i) s - - wkn-pres-index₁ : ∀ t i s → index₁ (wknStatementAt {Γ = Γ} t i s) ≡ index₁ s - wkn-pres-index₁ _ i (s ∙ s₁) = cong₂ (λ m n → suc (m ℕ.+ n)) (wkn-pres-index₁ _ i s) (wkn-pres-index₁ _ i s₁) - wkn-pres-index₁ _ i skip = refl - wkn-pres-index₁ _ i (ref ≔ x) = refl - wkn-pres-index₁ _ i (declare x s) = wkn-pres-index₁ _ (suc i) s - wkn-pres-index₁ _ i (invoke x x₁) = refl - wkn-pres-index₁ _ i (if x then s) = cong (λ m → suc (3 ℕ.* m)) (wkn-pres-index₁ _ i s) - wkn-pres-index₁ _ i (if x then s else s₁) = cong₂ (λ m n → suc (3 ℕ.* m ℕ.+ n)) (wkn-pres-index₁ _ i s) (wkn-pres-index₁ _ i s₁) - wkn-pres-index₁ _ i (for m s) = cong suc (wkn-pres-index₁ _ (suc i) s) - - wkn-pres-changes : ∀ t i {s} → ChangesState (wknStatementAt {Γ = Γ} t i s) → ChangesState s - wkn-pres-changes t i {_ ∙ _} (s ∙ˡ s₁) = wkn-pres-changes t i s ∙ˡ _ - wkn-pres-changes t i {_ ∙ _} (s ∙ʳ s₁) = _ ∙ʳ wkn-pres-changes t i s₁ - wkn-pres-changes t i {_ ≔ _} (_≔_ ref {canAssign} {refsState} e) = _≔_ _ {refsState = fromWitness (wknAt-pres-stateless i (toWitness refsState))} _ - wkn-pres-changes t i {declare _ _} (declare e s) = declare _ (wkn-pres-changes t (suc i) s) - wkn-pres-changes t i {invoke _ _} (invoke p es) = invoke _ _ - wkn-pres-changes t i {if _ then _} (if e then s) = if _ then wkn-pres-changes t i s - wkn-pres-changes t i {if _ then _ else _} (if e then′ s else s₁) = if _ then′ wkn-pres-changes t i s else _ - wkn-pres-changes t i {if _ then _ else _} (if e then s else′ s₁) = if _ then _ else′ wkn-pres-changes t i s₁ - wkn-pres-changes t i {for _ _} (for m s) = for m (wkn-pres-changes t (suc i) s) - - RecItem : Set - RecItem = ∃ λ n → ∃ (Statement {n}) - - inlinePredicate : RecItem → Set - inlinePredicate (_ , Γ , s) = ¬ ChangesState s → ∀ {ret} → (e : Expression Γ ret) → ∃ λ (e′ : Expression Γ ret) → callDepth e′ ℕ.≤ stmtCallDepth s ⊔ callDepth e - - inlineRel : RecItem → RecItem → Set - inlineRel = Lex.×-Lex _≡_ ℕ._<_ (Lex.×-Lex _≡_ ℕ._<_ ℕ._<_) on < (index₀ ∘ proj₂ ∘ proj₂) , < (index₁ ∘ proj₂ ∘ proj₂) , (index₂ ∘ proj₂ ∘ proj₂) > > - - inlineRelWf : Wf.WellFounded inlineRel - inlineRelWf = - On.wellFounded - < (index₀ ∘ proj₂ ∘ proj₂) , < (index₁ ∘ proj₂ ∘ proj₂) , (index₂ ∘ proj₂ ∘ proj₂) > > - (Lex.×-wellFounded ℕᵢ.<-wellFounded (Lex.×-wellFounded ℕᵢ.<-wellFounded ℕᵢ.<-wellFounded)) - - s<s∙s₁ : ∀ (s s₁ : Statement Γ) → inlineRel (_ , _ , s) (_ , _ , (s ∙ s₁)) - s<s∙s₁ s s₁ = case index₀ s₁ return (λ x → index₀ s ℕ.< index₀ s ℕ.+ x ⊎ index₀ s ≡ index₀ s ℕ.+ x × (index₁ s ℕ.< suc (index₁ s ℕ.+ index₁ s₁) ⊎ (index₁ s ≡ suc (index₁ s ℕ.+ index₁ s₁) × index₂ s ℕ.< 0))) of λ - { 0 → inj₂ (sym (ℕₚ.+-identityʳ (index₀ s)) , inj₁ (ℕₚ.m≤m+n (suc (index₁ s)) (index₁ s₁))) - ; (suc n) → inj₁ (ℕₚ.m<m+n (index₀ s) ℕₚ.0<1+n) - } - - s₁<s∙s₁ : ∀ (s s₁ : Statement Γ) → inlineRel (_ , _ , s₁) (_ , _ , (s ∙ s₁)) - s₁<s∙s₁ s s₁ = case index₀ s return (λ x → index₀ s₁ ℕ.< x ℕ.+ index₀ s₁ ⊎ index₀ s₁ ≡ x ℕ.+ index₀ s₁ × (index₁ s₁ ℕ.< suc (index₁ s ℕ.+ index₁ s₁) ⊎ (index₁ s₁ ≡ suc (index₁ s ℕ.+ index₁ s₁) × index₂ s₁ ℕ.< 0))) of λ - { 0 → inj₂ (refl , inj₁ (ℕ.s≤s (ℕₚ.m≤n+m (index₁ s₁) (index₁ s)))) - ; (suc n) → inj₁ (ℕₚ.m<n+m (index₀ s₁) ℕₚ.0<1+n) - } - s<declare‿s : ∀ (s : Statement _) (e : Expression Γ t) → inlineRel (_ , _ , s) (_ , _ , declare e s) - s<declare‿s s _ = inj₂ (refl , inj₂ (refl , ℕₚ.n<1+n (index₂ s))) - - splitIf : ∀ (x : Expression Γ bool) (s s₁ : Statement Γ) → Statement Γ - splitIf x s s₁ = declare x (if var 0F then wknStatementAt bool 0F s ∙ if var 0F then wknStatementAt bool 0F s₁) - - splitIf<if‿s∙s₁ : ∀ (x : Expression Γ bool) (s s₁ : Statement Γ) → inlineRel (_ , _ , splitIf x s s₁) (_ , _ , (if x then (s ∙ s₁))) - splitIf<if‿s∙s₁ x s s₁ = inj₂ (s≡₀s′ , inj₁ s<₁s′) - where - open +-*-Solver using (solve; _:*_; _:+_; con; _:=_) - s≡₀s′ = cong₂ ℕ._+_ (wkn-pres-index₀ bool 0F s) (wkn-pres-index₀ bool 0F s₁) - s<₁s′ = begin-strict - suc (suc (3 ℕ.* index₁ (wknStatementAt bool 0F s)) ℕ.+ suc (3 ℕ.* index₁ (wknStatementAt bool 0F s₁))) - ≡⟨ cong₂ (λ m n → suc (suc (3 ℕ.* m) ℕ.+ suc (3 ℕ.* n))) (wkn-pres-index₁ bool 0F s) (wkn-pres-index₁ bool 0F s₁) ⟩ - suc (suc (3 ℕ.* index₁ s) ℕ.+ suc (3 ℕ.* index₁ s₁)) - <⟨ ℕₚ.m<n+m (suc (suc (3 ℕ.* index₁ s) ℕ.+ suc (3 ℕ.* index₁ s₁))) (ℕₚ.0<1+n {n = 0}) ⟩ - suc (suc (suc (3 ℕ.* index₁ s) ℕ.+ suc (3 ℕ.* index₁ s₁))) - ≡⟨ solve 2 (λ m n → con 2 :+ (con 1 :+ (con 3 :* m) :+ (con 1 :+ (con 3 :* n))) := con 1 :+ (con 3 :* (con 1 :+ m :+ n))) refl (index₁ s) (index₁ s₁) ⟩ - suc (3 ℕ.* (suc (index₁ s ℕ.+ index₁ s₁))) +module Inline where + private + elses : LocalStatement Σ Γ → ℕ + elses (s ∙ s₁) = elses s ℕ.+ elses s₁ + elses skip = 0 + elses (ref ≔ val) = 0 + elses (declare e s) = elses s + elses (if x then s) = elses s + elses (if x then s else s₁) = 1 ℕ.+ elses s ℕ.+ elses s₁ + elses (for n s) = elses s + + structure : LocalStatement Σ Γ → ℕ + structure (s ∙ s₁) = 1 ℕ.+ structure s ℕ.+ structure s₁ + structure skip = 0 + structure (ref ≔ val) = 0 + structure (declare e s) = structure s + structure (if x then s) = 1 ℕ.+ 3 ℕ.* structure s + structure (if x then s else s₁) = 1 ℕ.+ 3 ℕ.* structure s ℕ.+ structure s₁ + structure (for n s) = 1 ℕ.+ structure s + + scope : LocalStatement Σ Γ → ℕ + scope (s ∙ s₁) = 0 + scope skip = 0 + scope (ref ≔ val) = 0 + scope (declare e s) = 1 ℕ.+ scope s + scope (if x then s) = 2 ℕ.* scope s + scope (if x then s else s₁) = 0 + scope (for n s) = 0 + + weaken-elses : ∀ i t′ (s : LocalStatement Σ Γ) → elses (Weaken.locStmt i t′ s) ≡ elses s + weaken-elses i t′ (s ∙ s₁) = cong₂ ℕ._+_ (weaken-elses i t′ s) (weaken-elses i t′ s₁) + weaken-elses i t′ skip = refl + weaken-elses i t′ (ref ≔ val) = refl + weaken-elses i t′ (declare e s) = weaken-elses (suc i) t′ s + weaken-elses i t′ (if x then s) = weaken-elses i t′ s + weaken-elses i t′ (if x then s else s₁) = cong₂ (λ m n → 1 ℕ.+ m ℕ.+ n) (weaken-elses i t′ s) (weaken-elses i t′ s₁) + weaken-elses i t′ (for n s) = weaken-elses (suc i) t′ s + + weaken-structure : ∀ i t′ (s : LocalStatement Σ Γ) → structure (Weaken.locStmt i t′ s) ≡ structure s + weaken-structure i t′ (s ∙ s₁) = cong₂ (λ m n → 1 ℕ.+ m ℕ.+ n) (weaken-structure i t′ s) (weaken-structure i t′ s₁) + weaken-structure i t′ skip = refl + weaken-structure i t′ (ref ≔ val) = refl + weaken-structure i t′ (declare e s) = weaken-structure (suc i) t′ s + weaken-structure i t′ (if x then s) = cong (λ m → 1 ℕ.+ 3 ℕ.* m) (weaken-structure i t′ s) + weaken-structure i t′ (if x then s else s₁) = cong₂ (λ m n → 1 ℕ.+ 3 ℕ.* m ℕ.+ n) (weaken-structure i t′ s) (weaken-structure i t′ s₁) + weaken-structure i t′ (for n s) = cong suc (weaken-structure (suc i) t′ s) + + weaken-scope : ∀ i t′ (s : LocalStatement Σ Γ) → scope (Weaken.locStmt i t′ s) ≡ scope s + weaken-scope i t′ (s ∙ s₁) = refl + weaken-scope i t′ skip = refl + weaken-scope i t′ (ref ≔ val) = refl + weaken-scope i t′ (declare e s) = cong suc (weaken-scope (suc i) t′ s) + weaken-scope i t′ (if x then s) = cong (2 ℕ.*_) (weaken-scope i t′ s) + weaken-scope i t′ (if x then s else s₁) = refl + weaken-scope i t′ (for n s) = refl + + RecItem : Vec Type n → Set + RecItem Σ = ∃ λ n → ∃ λ (Γ : Vec Type n) → LocalStatement Σ Γ + + P : ∀ (Σ : Vec Type n) → RecItem Σ → Set + P Σ (_ , Γ , s) = ∀ {t} (e : Expression Σ Γ t) → ∃ λ (e′ : Expression Σ Γ t) → CallDepth.expr e′ ≤ CallDepth.locStmt s ⊔ CallDepth.expr e + + index : RecItem Σ → ℕ × ℕ × ℕ + index = < elses , < structure , scope > > ∘ proj₂ ∘ proj₂ + + infix 4 _≺_ + + _≺_ : RecItem Σ → RecItem Σ → Set + _≺_ = ×-Lex _≡_ _<_ (×-Lex _≡_ _<_ _<_) on index + + ≺-wellFounded : WellFounded (_≺_ {Σ = Σ}) + ≺-wellFounded = On.wellFounded index (×-wellFounded ℕᵢ.<-wellFounded (×-wellFounded ℕᵢ.<-wellFounded ℕᵢ.<-wellFounded)) + + ≤∧<⇒≺ : ∀ (item item₁ : RecItem Σ) → (_≤_ on proj₁ ∘ index) item item₁ → (×-Lex _≡_ _<_ _<_ on proj₂ ∘ index) item item₁ → item ≺ item₁ + ≤∧<⇒≺ item item₁ ≤₁ <₂ with proj₁ (index item) ℕₚ.<? proj₁ (index item₁) + ... | yes <₁ = inj₁ <₁ + ... | no ≮₁ = inj₂ (ℕₚ.≤∧≮⇒≡ ≤₁ ≮₁ , <₂) + + pushIf : Expression Σ Γ bool → LocalStatement Σ Γ → LocalStatement Σ Γ + pushIf e′ (s ∙ s₁) = declare e′ (if var 0F then Weaken.locStmt 0F _ s ∙ if var 0F then Weaken.locStmt 0F _ s₁) + pushIf e′ skip = skip + pushIf e′ (ref ≔ val) = ref ≔ (if e′ then val else !! ref) + pushIf e′ (declare e s) = declare e (if Weaken.expr 0F _ e′ then s) + pushIf e′ (if x then s) = if e′ && x then s + pushIf e′ (if x then s else s₁) = declare (tup (e′ ∷ x ∷ [])) (if head (var 0F) && head (tail (var 0F)) then Weaken.locStmt 0F _ s ∙ if head (var 0F) && inv (head (tail (var 0F))) then Weaken.locStmt 0F _ s₁) + pushIf e′ (for n s) = declare e′ (for n (if var 1F then Weaken.locStmt 1F _ s)) + + pushIf≺if‿then : ∀ (e : Expression Σ Γ bool) s → (-, -, pushIf e s) ≺ (-, -, (if e then s)) + pushIf≺if‿then e′ (s ∙ s₁) = inj₂ + ( cong₂ ℕ._+_ (weaken-elses 0F bool s) (weaken-elses 0F bool s₁) + , inj₁ (begin-strict + 1 ℕ.+ (1 ℕ.+ 3 ℕ.* structure (Weaken.locStmt 0F bool s)) ℕ.+ (1 ℕ.+ 3 ℕ.* structure (Weaken.locStmt 0F bool s₁)) + ≡⟨ cong₂ (λ m n → 1 ℕ.+ (1 ℕ.+ 3 ℕ.* m) ℕ.+ (1 ℕ.+ 3 ℕ.* n)) (weaken-structure 0F bool s) (weaken-structure 0F bool s₁) ⟩ + 1 ℕ.+ (1 ℕ.+ 3 ℕ.* structure s) ℕ.+ (1 ℕ.+ 3 ℕ.* structure s₁) + <⟨ ℕₚ.n<1+n _ ⟩ + 2 ℕ.+ (1 ℕ.+ 3 ℕ.* structure s) ℕ.+ (1 ℕ.+ 3 ℕ.* structure s₁) + ≡⟨ solve-+ 2 + (λ a b → + con 2 :+ (con 1 :+ con 3 :* a) :+ (con 1 :+ con 3 :* b) := + con 1 :+ con 3 :* (con 1 :+ a :+ b)) + refl (structure s) (structure s₁) ⟩ + 1 ℕ.+ 3 ℕ.* (1 ℕ.+ structure s ℕ.+ structure s₁) + ∎) + ) + pushIf≺if‿then e′ skip = inj₂ (refl , inj₁ ℕₚ.0<1+n) + pushIf≺if‿then e′ (ref ≔ val) = inj₂ (refl , inj₁ ℕₚ.0<1+n) + pushIf≺if‿then e′ (declare e s) = inj₂ (refl , inj₂ (refl , (begin-strict + 1 ℕ.+ 2 ℕ.* scope s + <⟨ ℕₚ.n<1+n _ ⟩ + 2 ℕ.+ 2 ℕ.* scope s + ≡⟨ solve-+ 1 (λ a → con 2 :+ con 2 :* a := con 2 :* (con 1 :+ a)) refl (scope s) ⟩ + 2 ℕ.* (1 ℕ.+ scope s) + ∎))) + pushIf≺if‿then e′ (if x then s) = inj₂ (refl , (inj₁ (begin-strict + 1 ℕ.+ 3 ℕ.* structure s + <⟨ ℕₚ.m<n+m _ {3 ℕ.+ 6 ℕ.* structure s} ℕₚ.0<1+n ⟩ + 3 ℕ.+ 6 ℕ.* structure s ℕ.+ (1 ℕ.+ 3 ℕ.* structure s) + ≡⟨ solve-+ + 1 + (λ a → (con 3 :+ con 6 :* a) :+ (con 1 :+ con 3 :* a) + := con 1 :+ con 3 :* (con 1 :+ con 3 :* a)) + refl + (structure s) ⟩ + 1 ℕ.+ 3 ℕ.* (1 ℕ.+ 3 ℕ.* structure s) + ∎))) + pushIf≺if‿then e′ (if x then s else s₁) = inj₁ (begin-strict + elses (Weaken.locStmt 0F (tuple (bool ∷ bool ∷ [])) s) ℕ.+ elses (Weaken.locStmt 0F (tuple (bool ∷ bool ∷ [])) s₁) + ≡⟨ cong₂ ℕ._+_ (weaken-elses 0F _ s) (weaken-elses 0F _ s₁) ⟩ + elses s ℕ.+ elses s₁ + <⟨ ℕₚ.n<1+n _ ⟩ + 1 ℕ.+ elses s ℕ.+ elses s₁ + ∎) + pushIf≺if‿then e′ (for n s) = inj₂ + ( weaken-elses 1F bool s + , inj₁ (begin-strict + 2 ℕ.+ 3 ℕ.* structure (Weaken.locStmt 1F bool s) + ≡⟨ cong (λ m → 2 ℕ.+ 3 ℕ.* m) (weaken-structure 1F bool s) ⟩ + 2 ℕ.+ 3 ℕ.* structure s + <⟨ ℕₚ.m<n+m _ {2} ℕₚ.0<1+n ⟩ + 4 ℕ.+ 3 ℕ.* structure s + ≡⟨ solve-+ 1 (λ a → con 4 :+ con 3 :* a := con 1 :+ con 3 :* (con 1 :+ a)) refl (structure s) ⟩ + 1 ℕ.+ 3 ℕ.* (1 ℕ.+ structure s) + ∎) + ) + + pushIf-depth : ∀ (e : Expression Σ Γ bool) s → CallDepth.locStmt (pushIf e s) ≤ CallDepth.locStmt (if e then s) + pushIf-depth e′ (s ∙ s₁) = begin + CallDepth.locStmt (Weaken.locStmt 0F bool s) ⊔ 0 ⊔ (CallDepth.locStmt (Weaken.locStmt 0F bool s₁) ⊔ 0) ⊔ CallDepth.expr e′ + ≡⟨ cong₂ (λ m n → m ⊔ _ ⊔ (n ⊔ _) ⊔ _) (Weaken.locStmt-depth 0F bool s) (Weaken.locStmt-depth 0F bool s₁) ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ (CallDepth.locStmt s₁ ⊔ 0) ⊔ CallDepth.expr e′ + ≡⟨ solve-⊔ 3 (λ a b c → ((a ⊕ ε) ⊕ (b ⊕ ε)) ⊕ c ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locStmt s) _ _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.locStmt s₁ ⊔ CallDepth.expr e′ ∎ - - splitIf-stateless : ∀ {x : Expression Γ bool} {s s₁ : Statement Γ} → ¬ ChangesState (if x then (s ∙ s₁)) → ¬ ChangesState (splitIf x s s₁) - splitIf-stateless stateless (declare _ ((if _ then s) ∙ˡ _)) = stateless (if _ then (wkn-pres-changes bool 0F s ∙ˡ _)) - splitIf-stateless stateless (declare _ (_ ∙ʳ (if _ then s₁))) = stateless (if _ then (_ ∙ʳ wkn-pres-changes bool 0F s₁)) - - pushRef-stateless : ∀ {e} {ref : Expression Γ t} {canAssign val} → ¬ ChangesState (if e then _≔_ ref {canAssign} val) → ¬ ChangesState (_≔_ ref {canAssign} (if e then val else ref)) - pushRef-stateless stateless (_≔_ ref {refsState = refsState} _) = stateless (if _ then _≔_ ref {refsState = refsState} _) - - declare∘if<if∘declare : ∀ e (e′ : Expression Γ t) s → inlineRel (_ , _ , declare e′ (if wknAt 0F e then s)) (_ , _ , (if e then declare e′ s)) - declare∘if<if∘declare e e′ s = inj₂ (refl , inj₂ (refl , (begin-strict - suc (2 ℕ.* index₂ s) - <⟨ ℕₚ.n<1+n _ ⟩ - suc (suc (2 ℕ.* index₂ s)) - ≡⟨ solve 1 (λ m → con 2 :+ con 2 :* m := con 2 :* (con 1 :+ m)) refl (index₂ s) ⟩ - 2 ℕ.* suc (index₂ s) - ∎))) - where - open +-*-Solver using (solve; _:*_; _:+_; con; _:=_) - - declare∘if-stateless : ∀ {e} {e′ : Expression Γ t} {s} → ¬ ChangesState (if e then declare e′ s) → ¬ ChangesState (declare e′ (if wknAt 0F e then s)) - declare∘if-stateless stateless (declare _ (if _ then s)) = stateless (if _ then (declare _ s)) - - if<if∘if : ∀ (e e′ : Expression Γ bool) s → inlineRel (_ , _ , (if e && e′ then s)) (_ , _ , (if e then if e′ then s)) - if<if∘if e e′ s = inj₂ (refl , inj₁ (begin-strict - suc (3 ℕ.* index₁ s) - <⟨ ℕₚ.m<n+m (suc (3 ℕ.* index₁ s)) (ℕₚ.0<1+n {n = 2}) ⟩ - 4 ℕ.+ 3 ℕ.* index₁ s - ≤⟨ ℕₚ.m≤n+m (4 ℕ.+ 3 ℕ.* index₁ s) (6 ℕ.* index₁ s) ⟩ - 6 ℕ.* index₁ s ℕ.+ (4 ℕ.+ 3 ℕ.* index₁ s) - ≡⟨ solve 1 (λ m → con 6 :* m :+ (con 4 :+ con 3 :* m) := con 1 :+ con 3 :* (con 1 :+ con 3 :* m)) refl (index₁ s) ⟩ - suc (3 ℕ.* suc (3 ℕ.* index₁ s)) - ∎)) - where - open +-*-Solver using (solve; _:*_; _:+_; con; _:=_) - - if-stateless : ∀ {e e′ : Expression Γ bool} {s} → ¬ ChangesState (if e then if e′ then s) → ¬ ChangesState (if e && e′ then s) - if-stateless stateless (if _ then s) = stateless (if _ then if _ then s) - - if∙if : ∀ (e e′ : Expression Γ bool) (s s₁ : Statement Γ) → Statement Γ - if∙if e e′ s s₁ = - declare e ( - declare (wknAt 0F e′) ( - if var 1F && var 0F then wknStatementAt bool 0F (wknStatementAt bool 0F s) ∙ - if var 1F && inv (var 0F) then wknStatementAt bool 0F (wknStatementAt bool 0F s₁))) - - if∙if<if‿if‿else : ∀ (e e′ : Expression Γ bool) s s₁ → inlineRel (_ , _ , if∙if e e′ s s₁) (_ , _ , (if e then (if e′ then s else s₁))) - if∙if<if‿if‿else e e′ s s₁ = inj₁ (begin-strict - index₀ (wknStatementAt bool 0F (wknStatementAt bool 0F s)) ℕ.+ index₀ (wknStatementAt bool 0F (wknStatementAt bool 0F s₁)) - ≡⟨ cong₂ ℕ._+_ (wkn-pres-index₀ bool 0F (wknStatementAt bool 0F s)) (wkn-pres-index₀ bool 0F (wknStatementAt bool 0F s₁)) ⟩ - index₀ (wknStatementAt bool 0F s) ℕ.+ index₀ (wknStatementAt bool 0F s₁) - ≡⟨ cong₂ ℕ._+_ (wkn-pres-index₀ bool 0F s) (wkn-pres-index₀ bool 0F s₁) ⟩ - index₀ s ℕ.+ index₀ s₁ - <⟨ ℕₚ.n<1+n (index₀ s ℕ.+ index₀ s₁) ⟩ - suc (index₀ s ℕ.+ index₀ s₁) - ∎) - - if∙if-stateless : ∀ {e e′ : Expression Γ bool} {s s₁} → ¬ ChangesState (if e then (if e′ then s else s₁)) → ¬ ChangesState (if∙if e e′ s s₁) - if∙if-stateless stateless (declare _ (declare _ ((if _ then s) ∙ˡ _))) = stateless (if _ then (if _ then′ wkn-pres-changes bool 0F (wkn-pres-changes bool 0F s) else _)) - if∙if-stateless stateless (declare _ (declare _ (_ ∙ʳ (if _ then s₁)))) = stateless (if _ then (if _ then _ else′ wkn-pres-changes bool 0F (wkn-pres-changes bool 0F s₁))) - - declare-stateless : ∀ {i : Fin m} {s : Statement (fin m ∷ Γ)} → ¬ ChangesState (for m s) → ¬ ChangesState (declare (lit (i ′f)) s) - declare-stateless stateless (declare _ s) = stateless (for _ s) - - for‿if : ∀ (e : Expression Γ bool) m (s : Statement (fin m ∷ Γ)) → Statement Γ - for‿if e m s = declare e (for m (if var 1F then wknStatementAt bool 1F s)) - - for‿if<if‿for : ∀ (e : Expression Γ bool) m s → inlineRel (_ , _ , for‿if e m s) (_ , _ , (if e then for m s)) - for‿if<if‿for e m s = inj₂ (s≡₀s′ , inj₁ s<₁s′) - where - open +-*-Solver using (solve; _:*_; _:+_; con; _:=_) - s≡₀s′ = wkn-pres-index₀ bool 1F s - s<₁s′ = begin-strict - suc (suc (3 ℕ.* index₁ (wknStatementAt bool 1F s))) - ≡⟨ cong (λ m → suc (suc (3 ℕ.* m))) (wkn-pres-index₁ bool 1F s) ⟩ - suc (suc (3 ℕ.* index₁ s)) - <⟨ ℕₚ.m<n+m (suc (suc (3 ℕ.* index₁ s))) (ℕₚ.0<1+n {n = 1}) ⟩ - 4 ℕ.+ 3 ℕ.* index₁ s - ≡⟨ solve 1 (λ m → con 4 :+ con 3 :* m := con 1 :+ con 3 :* (con 1 :+ m)) refl (index₁ s) ⟩ - suc (3 ℕ.* suc (index₁ s)) + pushIf-depth e′ skip = z≤n + pushIf-depth e′ (ref ≔ val) = begin + CallDepth.locRef ref ⊔ (CallDepth.expr e′ ⊔ CallDepth.expr val ⊔ CallDepth.expr (!! ref)) + ≡⟨ cong (λ x → CallDepth.locRef ref ⊔ (CallDepth.expr e′ ⊔ CallDepth.expr val ⊔ x)) (CallDepth.homo-!! ref) ⟩ + CallDepth.locRef ref ⊔ (CallDepth.expr e′ ⊔ CallDepth.expr val ⊔ CallDepth.locRef ref) + ≡⟨ solve-⊔ 3 (λ a b c → a ⊕ ((c ⊕ b) ⊕ a) ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locRef ref) _ _ ⟩ + CallDepth.locRef ref ⊔ CallDepth.expr val ⊔ CallDepth.expr e′ + ∎ + pushIf-depth e′ (declare {t = t} e s) = begin + CallDepth.locStmt s ⊔ CallDepth.expr (Weaken.expr 0F t e′) ⊔ CallDepth.expr e + ≡⟨ cong (λ x → CallDepth.locStmt s ⊔ x ⊔ _) (Weaken.expr-depth 0F t e′) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e′ ⊔ CallDepth.expr e + ≡⟨ solve-⊔ 3 (λ a b c → (a ⊕ c) ⊕ b ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locStmt s) _ _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e ⊔ CallDepth.expr e′ + ∎ + pushIf-depth e′ (if x then s) = ℕₚ.≤-reflexive (solve-⊔ 3 (λ a b c → a ⊕ (c ⊕ b) ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locStmt s) _ _) + pushIf-depth e′ (if x then s else s₁) = begin + CallDepth.locStmt (Weaken.locStmt 0F (tuple (bool ∷ bool ∷ [])) s) ⊔ 0 ⊔ (CallDepth.locStmt (Weaken.locStmt 0F (tuple (bool ∷ bool ∷ [])) s₁) ⊔ 0) ⊔ (CallDepth.expr e′ ⊔ (CallDepth.expr x ⊔ 0)) + ≡⟨ cong₂ (λ m n → m ⊔ 0 ⊔ (n ⊔ 0) ⊔ _) (Weaken.locStmt-depth 0F _ s) (Weaken.locStmt-depth 0F _ s₁) ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ (CallDepth.locStmt s₁ ⊔ 0) ⊔ (CallDepth.expr e′ ⊔ (CallDepth.expr x ⊔ 0)) + ≡⟨ solve-⊔ 4 (λ a b c d → ((a ⊕ ε) ⊕ (b ⊕ ε)) ⊕ (d ⊕ (c ⊕ ε)) ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (CallDepth.locStmt s) _ _ _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.locStmt s₁ ⊔ CallDepth.expr x ⊔ CallDepth.expr e′ + ∎ + pushIf-depth e′ (for n s) = begin + CallDepth.locStmt (Weaken.locStmt 1F bool s) ⊔ 0 ⊔ CallDepth.expr e′ + ≡⟨ cong (λ x → x ⊔ _ ⊔ _) (Weaken.locStmt-depth 1F bool s) ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ CallDepth.expr e′ + ≡⟨ cong (_⊔ _) (ℕₚ.⊔-identityʳ (CallDepth.locStmt s)) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e′ ∎ - for‿if-stateless : ∀ {e : Expression Γ bool} {m s} → ¬ ChangesState (if e then for m s) → ¬ ChangesState (for‿if e m s) - for‿if-stateless stateless (declare _ (for _ (if _ then s))) = stateless (if _ then (for _ (wkn-pres-changes bool 1F s))) - - if∙if′ : ∀ (e : Expression Γ bool) (s s₁ : Statement Γ) → Statement Γ - if∙if′ e s s₁ = declare e ( - if var 0F then wknStatementAt bool 0F s ∙ - if inv (var 0F) then wknStatementAt bool 0F s₁) - - if∙if′<if‿else : ∀ (e : Expression Γ bool) s s₁ → inlineRel (_ , _ , if∙if′ e s s₁) (_ , _ , (if e then s else s₁)) - if∙if′<if‿else e s s₁ = inj₁ (begin-strict - index₀ (wknStatementAt bool 0F s) ℕ.+ index₀ (wknStatementAt bool 0F s₁) - ≡⟨ cong₂ ℕ._+_ (wkn-pres-index₀ bool 0F s) (wkn-pres-index₀ bool 0F s₁) ⟩ - index₀ s ℕ.+ index₀ s₁ - <⟨ ℕₚ.n<1+n (index₀ s ℕ.+ index₀ s₁) ⟩ - suc (index₀ s ℕ.+ index₀ s₁) - ∎) - - if∙if′-stateless : ∀ {e : Expression Γ bool} {s s₁} → ¬ ChangesState (if e then s else s₁) → ¬ ChangesState (if∙if′ e s s₁) - if∙if′-stateless stateless (declare _ ((if _ then s) ∙ˡ _)) = stateless (if _ then′ wkn-pres-changes bool 0F s else _) - if∙if′-stateless stateless (declare _ (_ ∙ʳ (if _ then s₁))) = stateless (if _ then _ else′ wkn-pres-changes bool 0F s₁) - - inlineHelper : ∀ n,Γ,s → Wf.WfRec inlineRel inlinePredicate n,Γ,s → inlinePredicate n,Γ,s - proj₁ (inlineHelper (_ , _ , (s ∙ s₁)) rec stateless e) = - proj₁ (rec - (_ , _ , s₁) - (s₁<s∙s₁ s s₁) - (stateless ∘ (s ∙ʳ_)) - (proj₁ (rec - (_ , _ , s) - (s<s∙s₁ s s₁) - (stateless ∘ (_∙ˡ s₁)) - e))) - proj₁ (inlineHelper (_ , _ , skip) rec stateless e) = e - proj₁ (inlineHelper (_ , _ , (_≔_ ref {canAssign} val)) rec stateless e) = - updateRef - (toWitness canAssign) - (stateless ∘ λ x → _≔_ ref {refsState = fromWitness x} val) - val - e - proj₁ (inlineHelper (_ , _ , declare x s) rec stateless e) = - elimAt 0F - (proj₁ (rec - (_ , _ , s) - (s<declare‿s s x) - (stateless ∘ declare x) - (wknAt 0F e))) - x - proj₁ (inlineHelper (_ , _ , invoke p es) rec stateless e) = contradiction (invoke p es) stateless - proj₁ (inlineHelper (_ , _ , (if x then (s ∙ s₁))) rec stateless e) = - proj₁ (rec - (_ , _ , splitIf x s s₁) - (splitIf<if‿s∙s₁ x s s₁) - (splitIf-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then skip)) rec stateless e) = e - proj₁ (inlineHelper (_ , _ , (if x then (_≔_ ref {canAssign} val))) rec stateless e) = - proj₁ (rec - (_ , _ , (_≔_ ref {canAssign} (if x then val else ref))) - (inj₂ (refl , inj₁ ℕₚ.0<1+n)) - (pushRef-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then declare x₁ s)) rec stateless e) = - proj₁ (rec - (_ , _ , declare x₁ (if wknAt 0F x then s)) - (declare∘if<if∘declare x x₁ s) - (declare∘if-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then invoke p es)) rec stateless e) = contradiction (if x then invoke p es) stateless - proj₁ (inlineHelper (_ , _ , (if x then if x₁ then s)) rec stateless e) = - proj₁ (rec - (_ , _ , (if (x && x₁) then s)) - (if<if∘if x x₁ s) - (if-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then (if x₁ then s else s₁))) rec stateless e) = - proj₁ (rec - (_ , _ , if∙if x x₁ s s₁) - (if∙if<if‿if‿else x x₁ s s₁) - (if∙if-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then for m s)) rec stateless e) = - proj₁ (rec - (_ , _ , for‿if x m s) - (for‿if<if‿for x m s) - (for‿if-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , (if x then s else s₁)) rec stateless e) = - proj₁ (rec - (_ , _ , if∙if′ x s s₁) - (if∙if′<if‿else x s s₁) - (if∙if′-stateless stateless) - e) - proj₁ (inlineHelper (_ , _ , for m s) rec stateless e) = - Vec.foldl - (λ _ → Expression _ _) - (λ e i → - proj₁ (rec - (_ , _ , declare (lit (i ′f)) s) - (inj₂ (refl , inj₁ (ℕₚ.n<1+n (index₁ s)))) - (declare-stateless stateless) - e)) + s≺s∙s₁ : ∀ (s s₁ : LocalStatement Σ Γ) → (-, -, s) ≺ (-, -, (s ∙ s₁)) + s≺s∙s₁ s s₁ = ≤∧<⇒≺ (-, -, s) (-, -, (s ∙ s₁)) (ℕₚ.m≤m+n (elses s) _) (inj₁ (ℕₚ.m≤m+n (suc (structure s)) _)) + + s₁≺s∙s₁ : ∀ (s s₁ : LocalStatement Σ Γ) → (-, -, s₁) ≺ (-, -, (s ∙ s₁)) + s₁≺s∙s₁ s s₁ = ≤∧<⇒≺ (-, -, s₁) (-, -, (s ∙ s₁)) (ℕₚ.m≤n+m _ (elses s)) (inj₁ (ℕₚ.m<n+m _ ℕₚ.0<1+n)) + + pushIfElse : Expression Σ Γ bool → LocalStatement Σ Γ → LocalStatement Σ Γ → LocalStatement Σ Γ + pushIfElse e s s₁ = declare e (if var 0F then Weaken.locStmt 0F _ s ∙ if inv (var 0F) then Weaken.locStmt 0F _ s₁) + + pushIfElse≺if‿then‿else : ∀ (e : Expression Σ Γ bool) s s₁ → (-, -, pushIfElse e s s₁) ≺ (-, -, (if e then s else s₁)) + pushIfElse≺if‿then‿else e s s₁ = inj₁ (begin-strict + elses (Weaken.locStmt 0F bool s) ℕ.+ elses (Weaken.locStmt 0F bool s₁) + ≡⟨ cong₂ ℕ._+_ (weaken-elses 0F bool s) (weaken-elses 0F bool s₁) ⟩ + elses s ℕ.+ elses s₁ + <⟨ ℕₚ.n<1+n _ ⟩ + 1 ℕ.+ elses s ℕ.+ elses s₁ + ∎) + + helper : ∀ item → Wf.WfRec _≺_ (P Σ) item → P Σ item + proj₁ (helper (_ , _ , (s ∙ s₁)) rec e) = proj₁ (rec (-, -, s) (s≺s∙s₁ s s₁) (proj₁ (rec (-, -, s₁) (s₁≺s∙s₁ s s₁) e))) + proj₁ (helper (_ , _ , skip) rec e) = e + proj₁ (helper (_ , _ , (ref ≔ val)) rec e) = Update.expr ref val e + proj₁ (helper (_ , _ , declare x s) rec e) = Elim.expr 0F (proj₁ (rec (-, -, s) (inj₂ (refl , inj₂ (refl , ℕₚ.n<1+n _))) (Weaken.expr 0F _ e))) x + proj₁ (helper (_ , _ , (if x then s)) rec e) = proj₁ (rec (-, -, pushIf x s) (pushIf≺if‿then x s) e) + proj₁ (helper (_ , _ , (if x then s else s₁)) rec e) = proj₁ (rec (-, -, pushIfElse x s s₁) (pushIfElse≺if‿then‿else x s s₁) e) + proj₁ (helper (_ , _ , for n s) rec e) = Vec.foldr (λ _ → Expression _ _ _) (λ i e → proj₁ (rec (-, -, declare (lit i) s) (inj₂ (refl , inj₁ (ℕₚ.n<1+n (structure s)))) e)) e (Vec.allFin n) + proj₂ (helper (_ , _ , (s ∙ s₁)) rec e) = begin + CallDepth.expr (proj₁ outer) + ≤⟨ proj₂ outer ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr (proj₁ inner) + ≤⟨ ℕₚ.⊔-monoʳ-≤ (CallDepth.locStmt s) (proj₂ inner) ⟩ + CallDepth.locStmt s ⊔ (CallDepth.locStmt s₁ ⊔ CallDepth.expr e) + ≡˘⟨ ℕₚ.⊔-assoc (CallDepth.locStmt s) _ _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.locStmt s₁ ⊔ CallDepth.expr e + ∎ + where + inner = rec (-, -, s₁) (s₁≺s∙s₁ s s₁) e + outer = rec (-, -, s) (s≺s∙s₁ s s₁) (proj₁ inner) + proj₂ (helper (_ , _ , skip) rec e) = ℕₚ.≤-refl + proj₂ (helper (_ , _ , (ref ≔ val)) rec e) = begin + CallDepth.expr (Update.expr ref val e) + ≤⟨ Update.expr-depth ref val e ⟩ + CallDepth.expr e ⊔ (CallDepth.locRef ref ⊔ CallDepth.expr val) + ≡⟨ ℕₚ.⊔-comm (CallDepth.expr e) _ ⟩ + CallDepth.locRef ref ⊔ CallDepth.expr val ⊔ CallDepth.expr e + ∎ + proj₂ (helper (_ , _ , declare x s) rec e) = begin + CallDepth.expr (Elim.expr 0F (proj₁ inner) x) + ≤⟨ Elim.expr-depth 0F (proj₁ inner) x ⟩ + CallDepth.expr (proj₁ inner) ⊔ CallDepth.expr x + ≤⟨ ℕₚ.⊔-monoˡ-≤ _ (proj₂ inner) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr (Weaken.expr 0F _ e) ⊔ CallDepth.expr x + ≡⟨ cong (λ x → CallDepth.locStmt s ⊔ x ⊔ _) (Weaken.expr-depth 0F _ e) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e ⊔ CallDepth.expr x + ≡⟨ solve-⊔ 3 (λ a b c → (a ⊕ c) ⊕ b ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.locStmt s) _ _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr x ⊔ CallDepth.expr e + ∎ + where inner = rec (-, -, s) _ (Weaken.expr 0F _ e) + proj₂ (helper (_ , _ , (if x then s)) rec e) = begin + CallDepth.expr (proj₁ inner) + ≤⟨ proj₂ inner ⟩ + CallDepth.locStmt (pushIf x s) ⊔ CallDepth.expr e + ≤⟨ ℕₚ.⊔-monoˡ-≤ _ (pushIf-depth x s) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr x ⊔ CallDepth.expr e + ∎ + where inner = rec (-, -, pushIf x s) (pushIf≺if‿then x s) e + proj₂ (helper (_ , _ , (if x then s else s₁)) rec e) = begin + CallDepth.expr (proj₁ inner) + ≤⟨ proj₂ inner ⟩ + CallDepth.locStmt (Weaken.locStmt 0F bool s) ⊔ 0 ⊔ (CallDepth.locStmt (Weaken.locStmt 0F bool s₁) ⊔ 0) ⊔ CallDepth.expr x ⊔ CallDepth.expr e + ≡⟨ cong₂ (λ m n → m ⊔ 0 ⊔ (n ⊔ 0) ⊔ _ ⊔ _) (Weaken.locStmt-depth 0F bool s) (Weaken.locStmt-depth 0F bool s₁) ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ (CallDepth.locStmt s₁ ⊔ 0) ⊔ CallDepth.expr x ⊔ CallDepth.expr e + ≡⟨ cong (λ x → x ⊔ _ ⊔ _) (solve-⊔ 2 (λ a b → (a ⊕ ε) ⊕ (b ⊕ ε) ⊜ a ⊕ b) refl (CallDepth.locStmt s) _) ⟩ + CallDepth.locStmt s ⊔ CallDepth.locStmt s₁ ⊔ CallDepth.expr x ⊔ CallDepth.expr e + ∎ + where inner = rec (-, -, pushIfElse x s s₁) (pushIfElse≺if‿then‿else x s s₁) e + proj₂ (helper (_ , _ , for n s) rec e) = foldr-lubs + _ e - (Vec.allFin _) - proj₂ (inlineHelper (_ , _ , (s ∙ s₁)) rec stateless e) = begin - callDepth (proj₁ outer) - ≤⟨ proj₂ outer ⟩ - stmtCallDepth s₁ ⊔ callDepth (proj₁ inner) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (stmtCallDepth s₁) (proj₂ inner) ⟩ - stmtCallDepth s₁ ⊔ (stmtCallDepth s ⊔ callDepth e) - ≡⟨ ⊔-solve 3 (λ a b c → b ⊕ (a ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (stmtCallDepth s) (stmtCallDepth s₁) (callDepth e) ⟩ - stmtCallDepth s ⊔ stmtCallDepth s₁ ⊔ callDepth e + (λ e′ → CallDepth.expr e′ ≤ CallDepth.locStmt s ⊔ CallDepth.expr e) + (ℕₚ.m≤n⊔m (CallDepth.locStmt s) _) + (λ i {e′} e′≤s⊔e → begin + CallDepth.expr (proj₁ (rec (-, -, declare (lit i) s) _ e′)) + ≤⟨ proj₂ (rec (-, -, declare (lit i) s) _ e′) ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ CallDepth.expr e′ + ≤⟨ ℕₚ.⊔-monoʳ-≤ (CallDepth.locStmt s ⊔ _) e′≤s⊔e ⟩ + CallDepth.locStmt s ⊔ 0 ⊔ (CallDepth.locStmt s ⊔ CallDepth.expr e) + ≡⟨ solve-⊔ 2 (λ a b → (a ⊕ ε) ⊕ (a ⊕ b) ⊜ a ⊕ b) refl (CallDepth.locStmt s) _ ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e + ∎) + (Vec.allFin n) + + helper′ : ∀ (s : LocalStatement Σ Γ) (e : Expression Σ Γ t) → ∃ λ (e′ : Expression Σ Γ t) → CallDepth.expr e′ ≤ CallDepth.locStmt s ⊔ CallDepth.expr e + helper′ s = Wf.All.wfRec ≺-wellFounded _ (P _) helper (-, -, s) + + fun : Function Σ Δ ret → All (Expression Σ Γ) Δ → Expression Σ Γ ret + fun (declare e f) es = fun f (SubstAll.expr e es ∷ es) + fun (s ∙return e) es = SubstAll.expr (proj₁ (helper′ s e)) es + + fun-depth : ∀ (f : Function Σ Δ ret) (es : All (Expression Σ Γ) Δ) → CallDepth.expr (fun f es) ≤ CallDepth.fun f ⊔ CallDepth.exprs es + fun-depth (declare e f) es = begin + CallDepth.expr (fun f (SubstAll.expr e es ∷ es)) + ≤⟨ fun-depth f (SubstAll.expr e es ∷ es) ⟩ + CallDepth.fun f ⊔ (CallDepth.exprs es ⊔ CallDepth.expr (SubstAll.expr e es)) + ≤⟨ ℕₚ.⊔-monoʳ-≤ (CallDepth.fun f) (ℕₚ.⊔-monoʳ-≤ (CallDepth.exprs es) (SubstAll.expr-depth e es)) ⟩ + CallDepth.fun f ⊔ (CallDepth.exprs es ⊔ (CallDepth.expr e ⊔ CallDepth.exprs es)) + ≡⟨ solve-⊔ 3 (λ a b c → a ⊕ (c ⊕ (b ⊕ c)) ⊜ (a ⊕ b) ⊕ c) refl (CallDepth.fun f) _ _ ⟩ + CallDepth.fun f ⊔ CallDepth.expr e ⊔ CallDepth.exprs es ∎ - where - inner = rec (_ , _ , s) (s<s∙s₁ s s₁) (stateless ∘ (_∙ˡ s₁)) e - outer = rec (_ , _ , s₁) (s₁<s∙s₁ s s₁) (stateless ∘ (s ∙ʳ_)) (proj₁ inner) - -- with rec (_ , _ , s) (s<s∙s₁ s s₁) (stateless ∘ (_∙ˡ s₁)) e - -- ... | inner , eq with inner | eq | rec (_ , _ , s₁) (s₁<s∙s₁ s s₁) (stateless ∘ (s ∙ʳ_)) inner - -- ... | inner | inj₁ inner≤s | outer , inj₁ outer≤s₁ = inj₁ (ℕₚ.m≤n⇒m≤o⊔n (stmtCallDepth s) outer≤s₁) - -- ... | inner | inj₁ inner≤s | outer , inj₂ outer≡inner = inj₁ (begin - -- callDepth outer ≡⟨ outer≡inner ⟩ - -- callDepth inner ≤⟨ ℕₚ.m≤n⇒m≤n⊔o (stmtCallDepth s₁) inner≤s ⟩ - -- stmtCallDepth s ⊔ stmtCallDepth s₁ ∎) - -- ... | inner | inj₂ inner≡e | outer , inj₁ outer≤s₁ = inj₁ (ℕₚ.m≤n⇒m≤o⊔n (stmtCallDepth s) outer≤s₁) - -- ... | inner | inj₂ inner≡e | outer , inj₂ outer≡inner = inj₂ (trans outer≡inner inner≡e) - proj₂ (inlineHelper (_ , _ , skip) rec stateless e) = ℕₚ.≤-refl - proj₂ (inlineHelper (_ , _ , (_≔_ ref {canAssign} x)) rec stateless e) = updateRef-pres-callDepth (toWitness canAssign) (λ x → stateless (_≔_ _ {refsState = fromWitness x} _)) x e - proj₂ (inlineHelper (_ , _ , declare x s) rec stateless e) = begin - callDepth (elimAt 0F (proj₁ inner) x) - ≤⟨ elimAt-pres-callDepth 0F (proj₁ inner) x ⟩ - callDepth x ⊔ callDepth (proj₁ inner) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (callDepth x) (proj₂ inner) ⟩ - callDepth x ⊔ (stmtCallDepth s ⊔ callDepth (wknAt 0F e)) - ≡⟨ cong (λ m → callDepth x ⊔ (stmtCallDepth s ⊔ m)) (wknAt-pres-callDepth 0F e) ⟩ - callDepth x ⊔ (stmtCallDepth s ⊔ callDepth e) - ≡⟨ ⊔-solve 3 (λ a b c → a ⊕ (b ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (callDepth x) (stmtCallDepth s) (callDepth e) ⟩ - callDepth x ⊔ stmtCallDepth s ⊔ callDepth e + fun-depth (s ∙return e) es = begin + CallDepth.expr (SubstAll.expr (proj₁ (helper′ s e)) es) + ≤⟨ SubstAll.expr-depth (proj₁ (helper′ s e)) es ⟩ + CallDepth.expr (proj₁ (helper′ s e)) ⊔ CallDepth.exprs es + ≤⟨ ℕₚ.⊔-monoˡ-≤ _ (proj₂ (helper′ s e)) ⟩ + CallDepth.locStmt s ⊔ CallDepth.expr e ⊔ CallDepth.exprs es ∎ - where - inner = rec (_ , _ , s) (s<declare‿s s x) (λ x → stateless (declare _ x)) (wknAt 0F e) - proj₂ (inlineHelper (_ , _ , invoke p es) rec stateless e) = contradiction (invoke p es) stateless - proj₂ (inlineHelper (_ , _ , (if x then (s ∙ s₁))) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x ⊔ (stmtCallDepth (wknStatementAt bool 0F s) ⊔ stmtCallDepth (wknStatementAt bool 0F s₁)) ⊔ callDepth e - ≡⟨ cong₂ (λ m n → callDepth x ⊔ (m ⊔ n) ⊔ callDepth e) (wkn-pres-callDepth bool 0F s) (wkn-pres-callDepth bool 0F s₁) ⟩ - callDepth x ⊔ (stmtCallDepth s ⊔ stmtCallDepth s₁) ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , splitIf x s s₁) (splitIf<if‿s∙s₁ x s s₁) (splitIf-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then skip)) rec stateless e) = ℕₚ.m≤n⊔m (callDepth x ⊔ 0) (callDepth e) - proj₂ (inlineHelper (_ , _ , (if x then (_≔_ ref {canAssign} val))) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth ref ⊔ (callDepth x ⊔ callDepth val ⊔ callDepth ref) ⊔ callDepth e - ≡⟨ ⊔-solve 4 (λ a b c d → (b ⊕ ((a ⊕ c) ⊕ b)) ⊕ d ⊜ (a ⊕ (b ⊕ c)) ⊕ d) refl (callDepth x) (callDepth ref) (callDepth val) (callDepth e) ⟩ - callDepth x ⊔ (callDepth ref ⊔ callDepth val) ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , (_≔_ ref {canAssign} (if x then val else ref))) (inj₂ (refl , inj₁ ℕₚ.0<1+n)) (pushRef-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then declare x₁ s)) rec stateless e) = begin - callDepth(proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x₁ ⊔ (callDepth (wknAt 0F x) ⊔ stmtCallDepth s) ⊔ callDepth e - ≡⟨ cong (λ m → callDepth x₁ ⊔ (m ⊔ stmtCallDepth s) ⊔ callDepth e) (wknAt-pres-callDepth 0F x) ⟩ - callDepth x₁ ⊔ (callDepth x ⊔ stmtCallDepth s) ⊔ callDepth e - ≡⟨ ⊔-solve 4 (λ a b c d → (b ⊕ (a ⊕ c)) ⊕ d ⊜ (a ⊕ (b ⊕ c)) ⊕ d) refl (callDepth x) (callDepth x₁) (stmtCallDepth s) (callDepth e) ⟩ - callDepth x ⊔ (callDepth x₁ ⊔ stmtCallDepth s) ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , declare x₁ (if wknAt 0F x then s)) (declare∘if<if∘declare x x₁ s) (declare∘if-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then invoke p es)) rec stateless e) = contradiction (if _ then invoke p es) stateless - proj₂ (inlineHelper (_ , _ , (if x then (if x₁ then s))) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x ⊔ callDepth x₁ ⊔ stmtCallDepth s ⊔ callDepth e - ≡⟨ ⊔-solve 4 (λ a b c d → ((a ⊕ b) ⊕ c) ⊕ d ⊜ (a ⊕ (b ⊕ c)) ⊕ d) refl (callDepth x) (callDepth x₁) (stmtCallDepth s) (callDepth e) ⟩ - callDepth x ⊔ (callDepth x₁ ⊔ stmtCallDepth s) ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , (if x && x₁ then s)) (if<if∘if x x₁ s) (if-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then (if x₁ then s else s₁))) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x ⊔ (callDepth (wknAt 0F x₁) ⊔ (stmtCallDepth (wknStatementAt bool 0F (wknStatementAt bool 0F s)) ⊔ stmtCallDepth (wknStatementAt bool 0F (wknStatementAt bool 0F s₁)))) ⊔ callDepth e - ≡⟨ congₙ 3 (λ m n o → callDepth x ⊔ (m ⊔ (n ⊔ o)) ⊔ callDepth e) (wknAt-pres-callDepth 0F x₁) (trans (wkn-pres-callDepth bool 0F (wknStatementAt bool 0F s)) (wkn-pres-callDepth bool 0F s)) (trans (wkn-pres-callDepth bool 0F (wknStatementAt bool 0F s₁)) (wkn-pres-callDepth bool 0F s₁)) ⟩ - callDepth x ⊔ (callDepth x₁ ⊔ (stmtCallDepth s ⊔ stmtCallDepth s₁)) ⊔ callDepth e - ≡⟨ ⊔-solve 5 (λ a b c d e → (a ⊕ (b ⊕ (c ⊕ d))) ⊕ e ⊜ (a ⊕ ((b ⊕ c) ⊕ d)) ⊕ e) refl (callDepth x) (callDepth x₁) (stmtCallDepth s) (stmtCallDepth s₁) (callDepth e) ⟩ - callDepth x ⊔ (callDepth x₁ ⊔ stmtCallDepth s ⊔ stmtCallDepth s₁) ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , if∙if x x₁ s s₁) (if∙if<if‿if‿else x x₁ s s₁) (if∙if-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then for m s)) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x ⊔ stmtCallDepth (wknStatementAt bool 1F s) ⊔ callDepth e - ≡⟨ cong (λ m → callDepth x ⊔ m ⊔ callDepth e) (wkn-pres-callDepth bool 1F s) ⟩ - callDepth x ⊔ stmtCallDepth s ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , for‿if x m s) (for‿if<if‿for x m s) (for‿if-stateless stateless) e - proj₂ (inlineHelper (_ , _ , (if x then s else s₁)) rec stateless e) = begin - callDepth (proj₁ inner) - ≤⟨ proj₂ inner ⟩ - callDepth x ⊔ (stmtCallDepth (wknStatementAt bool 0F s) ⊔ stmtCallDepth (wknStatementAt bool 0F s₁)) ⊔ callDepth e - ≡⟨ cong₂ (λ m n → callDepth x ⊔ (m ⊔ n) ⊔ callDepth e) (wkn-pres-callDepth bool 0F s) (wkn-pres-callDepth bool 0F s₁) ⟩ - callDepth x ⊔ (stmtCallDepth s ⊔ stmtCallDepth s₁) ⊔ callDepth e - ≡⟨ ⊔-solve 4 (λ a b c d → (a ⊕ (b ⊕ c)) ⊕ d ⊜ ((a ⊕ b) ⊕ c) ⊕ d) refl (callDepth x) (stmtCallDepth s) (stmtCallDepth s₁) (callDepth e) ⟩ - callDepth x ⊔ stmtCallDepth s ⊔ stmtCallDepth s₁ ⊔ callDepth e - ∎ - where - inner = rec (_ , _ , if∙if′ x s s₁) (if∙if′<if‿else x s s₁) (if∙if′-stateless stateless) e - proj₂ (inlineHelper (n , Γ , for m s) rec stateless {ret} e) = helper - (stmtCallDepth s) - (λ e i → - rec - (_ , _ , declare (lit (i ′f)) s) - (inj₂ (refl , inj₁ (ℕₚ.n<1+n (index₁ s)))) - (declare-stateless stateless) - e) - e - (Vec.allFin m) - where - helper : ∀ {n m} k (f : ∀ {n : ℕ} e (i : Fin m) → ∃ λ e′ → callDepth e′ ℕ.≤ k ⊔ callDepth e) → ∀ e xs → callDepth (Vec.foldl (λ _ → Expression Γ ret) {n} (λ {n} e i → proj₁ (f {n} e i)) e xs) ℕ.≤ k ⊔ callDepth e - helper k f e [] = ℕₚ.m≤n⊔m k (callDepth e) - helper k f e (x ∷ xs) = begin - callDepth (Vec.foldl _ (λ e i → proj₁ (f e i)) (proj₁ (f e x)) xs) - ≤⟨ helper k f (proj₁ (f e x)) xs ⟩ - k ⊔ callDepth (proj₁ (f e x)) - ≤⟨ ℕₚ.⊔-monoʳ-≤ k (proj₂ (f e x)) ⟩ - k ⊔ (k ⊔ callDepth e) - ≡⟨ ⊔-solve 2 (λ a b → a ⊕ (a ⊕ b) ⊜ a ⊕ b) refl k (callDepth e) ⟩ - k ⊔ callDepth e - ∎ -inlineFunction : Function Γ ret → All (Expression Δ) Γ → Expression Δ ret -inlineFunction (declare e f) args = inlineFunction f (subst e args ∷ args) -inlineFunction (_∙return_ s {stateless} e) args = - subst - (proj₁ (Wf.All.wfRec - inlineRelWf - _ - inlinePredicate - inlineHelper - (_ , _ , s) - (toWitnessFalse stateless) - e)) - args - -inlineFunction-lowers-callDepth : ∀ (f : Function Δ ret) (args : All (Expression Γ) Δ) → callDepth (inlineFunction f args) ℕ.≤ funCallDepth f ⊔ callDepth′ args -inlineFunction-lowers-callDepth (declare e f) args = begin - callDepth (inlineFunction f (subst e args ∷ args)) - ≤⟨ inlineFunction-lowers-callDepth f (subst e args ∷ args) ⟩ - funCallDepth f ⊔ (callDepth (subst e args) ⊔ callDepth′ args) - ≤⟨ ℕₚ.⊔-monoʳ-≤ (funCallDepth f) (ℕₚ.⊔-monoˡ-≤ (callDepth′ args) (subst-pres-callDepth e args)) ⟩ - funCallDepth f ⊔ (callDepth e ⊔ callDepth′ args ⊔ callDepth′ args) - ≡⟨ ⊔-solve 3 (λ a b c → a ⊕ ((b ⊕ c) ⊕ c) ⊜ (a ⊕ b) ⊕ c) refl (funCallDepth f) (callDepth e) (callDepth′ args) ⟩ - funCallDepth f ⊔ callDepth e ⊔ callDepth′ args - ∎ -inlineFunction-lowers-callDepth (_∙return_ s {stateless} e) args = begin - callDepth (subst (proj₁ theCall) args) ≤⟨ subst-pres-callDepth (proj₁ theCall) args ⟩ - callDepth (proj₁ theCall) ⊔ callDepth′ args ≤⟨ ℕₚ.⊔-monoˡ-≤ (callDepth′ args) (proj₂ theCall) ⟩ - stmtCallDepth s ⊔ callDepth e ⊔ callDepth′ args ∎ - where - theCall = Wf.All.wfRec - inlineRelWf - _ - inlinePredicate - inlineHelper - (_ , _ , s) - (toWitnessFalse stateless) - e - -elimFunctions : Expression Γ t → ∃ λ (e : Expression Γ t) → callDepth e ≡ 0 -elimFunctions {Γ = Γ} = Wf.All.wfRec relWf _ pred helper ∘ (_ ,_) - where - index : Expression Γ t → ℕ - index′ : All (Expression Γ) Δ → ℕ - - index (Code.lit x) = 0 - index (Code.state i) = 0 - index (Code.var i) = 0 - index (Code.abort e) = suc (index e) - index (e Code.≟ e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (e Code.<? e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (Code.inv e) = suc (index e) - index (e Code.&& e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (e Code.|| e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (Code.not e) = suc (index e) - index (e Code.and e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (e Code.or e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index Code.[ e ] = suc (index e) - index (Code.unbox e) = suc (index e) - index (Code.splice e e₁ e₂) = 1 ℕ.+ index e ℕ.+ index e₁ ℕ.+ index e₂ - index (Code.cut e e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (Code.cast eq e) = suc (index e) - index (Code.- e) = suc (index e) - index (e Code.+ e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (e Code.* e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (e Code.^ x) = suc (index e) - index (e Code.>> x) = suc (index e) - index (Code.rnd e) = suc (index e) - index (Code.fin x e) = suc (index e) - index (Code.asInt e) = suc (index e) - index Code.nil = 0 - index (Code.cons e e₁) = 1 ℕ.+ index e ℕ.+ index e₁ - index (Code.head e) = suc (index e) - index (Code.tail e) = suc (index e) - index (Code.call f es) = index′ es - index (Code.if e then e₁ else e₂) = 1 ℕ.+ index e ℕ.+ index e₁ ℕ.+ index e₂ - - index′ [] = 1 - index′ (e ∷ es) = index e ℕ.+ index′ es - - pred : ∃ (Expression Γ) → Set - pred (t , e) = ∃ λ (e : Expression Γ t) → callDepth e ≡ 0 - - pred′ : All (Expression Γ) Δ → Set - pred′ {Δ = Δ} _ = ∃ λ (es : All (Expression Γ) Δ) → callDepth′ es ≡ 0 - - rel = Lex.×-Lex _≡_ ℕ._<_ ℕ._<_ on < callDepth ∘ proj₂ , index ∘ proj₂ > - - data _<′₁_ : ∃ (Expression Γ) → All (Expression Γ) Δ → Set where - inj₁ : {e : Expression Γ t} {es : All (Expression Γ) Δ} → callDepth e ℕ.< callDepth′ es → (t , e) <′₁ es - inj₂ : {e : Expression Γ t} {es : All (Expression Γ) Δ} → callDepth e ≡ callDepth′ es → index e ℕ.< index′ es → (t , e) <′₁ es - - data _<′₂_ : ∀ {n ts} → All (Expression Γ) {n} ts → All (Expression Γ) Δ → Set where - inj₁ : ∀ {n ts} {es′ : All (Expression Γ) {n} ts} {es : All (Expression Γ) Δ} → callDepth′ es′ ℕ.< callDepth′ es → es′ <′₂ es - inj₂ : ∀ {n ts} {es′ : All (Expression Γ) {n} ts} {es : All (Expression Γ) Δ} → callDepth′ es′ ≡ callDepth′ es → index′ es′ ℕ.≤ index′ es → es′ <′₂ es - - <′₁-<′₂-trans : ∀ {n ts} {e : Expression Γ t} {es : All (Expression Γ) Δ} {es′ : All (Expression Γ) {n} ts} → - (t , e) <′₁ es → es <′₂ es′ → (t , e) <′₁ es′ - <′₁-<′₂-trans (inj₁ <₁) (inj₁ <₂) = inj₁ (ℕₚ.<-trans <₁ <₂) - <′₁-<′₂-trans (inj₁ <₁) (inj₂ ≡₂ _) = inj₁ (ℕₚ.<-transˡ <₁ (ℕₚ.≤-reflexive ≡₂)) - <′₁-<′₂-trans (inj₂ ≡₁ _) (inj₁ <₂) = inj₁ (ℕₚ.<-transʳ (ℕₚ.≤-reflexive ≡₁) <₂) - <′₁-<′₂-trans (inj₂ ≡₁ <₁) (inj₂ ≡₂ ≤₂) = inj₂ (trans ≡₁ ≡₂) (ℕₚ.<-transˡ <₁ ≤₂) - - <′₂-trans : ∀ {m n ts ts′} {es : All (Expression Γ) Δ} {es′ : All (Expression Γ) {n} ts} {es′′ : All (Expression Γ) {m} ts′} → - es <′₂ es′ → es′ <′₂ es′′ → es <′₂ es′′ - <′₂-trans (inj₁ <₁) (inj₁ <₂) = inj₁ (ℕₚ.<-trans <₁ <₂) - <′₂-trans (inj₁ <₁) (inj₂ ≡₂ _) = inj₁ (ℕₚ.<-transˡ <₁ (ℕₚ.≤-reflexive ≡₂)) - <′₂-trans (inj₂ ≡₁ _) (inj₁ <₂) = inj₁ (ℕₚ.<-transʳ (ℕₚ.≤-reflexive ≡₁) <₂) - <′₂-trans (inj₂ ≡₁ ≤₁) (inj₂ ≡₂ ≤₂) = inj₂ (trans ≡₁ ≡₂) (ℕₚ.≤-trans ≤₁ ≤₂) - - index′>0 : ∀ (es : All (Expression Γ) Δ) → index′ es ℕ.> 0 - index′>0 [] = ℕₚ.≤-refl - index′>0 (e ∷ es) = ℕₚ.<-transˡ (index′>0 es) (ℕₚ.m≤n+m (index′ es) (index e)) - - e<′₁es⇒e<f[es] : ∀ {e : Expression Γ t} {es : All (Expression Γ) Δ} → (t , e) <′₁ es → ∀ f → rel (t , e) (t′ , call f es) - e<′₁es⇒e<f[es] {e = e} {es} (inj₁ <) f = inj₁ (ℕₚ.<-transˡ < (ℕₚ.m≤n⊔m (suc (funCallDepth f)) (callDepth′ es))) - e<′₁es⇒e<f[es] {e = e} {es} (inj₂ ≡ <) f with callDepth′ es | funCallDepth f | ℕ.compare (callDepth′ es) (suc (funCallDepth f)) - ... | _ | _ | ℕ.less m k = inj₁ (begin - suc (callDepth e) ≡⟨ cong suc ≡ ⟩ - suc m ≤⟨ ℕₚ.m≤m+n (suc m) k ⟩ - suc m ℕ.+ k ≤⟨ ℕₚ.m≤m⊔n (suc m ℕ.+ k) m ⟩ - suc m ℕ.+ k ⊔ m ∎) - ... | _ | _ | ℕ.equal m = inj₂ (trans ≡ (sym (ℕₚ.⊔-idem m)) , <) - ... | _ | _ | ℕ.greater n k = inj₂ ((begin-equality - callDepth e ≡⟨ ≡ ⟩ - suc n ℕ.+ k ≡˘⟨ ℕₚ.m≤n⇒m⊔n≡n (ℕₚ.≤-trans (ℕₚ.n≤1+n n) (ℕₚ.m≤m+n (suc n) k)) ⟩ - n ⊔ (suc n ℕ.+ k) ∎) , <) - - e<ᵢe∷es : ∀ (e : Expression Γ t) (es : All (Expression Γ) Δ) → index e ℕ.< index′ (e ∷ es) - e<ᵢe∷es e es = ℕₚ.m<m+n (index e) (index′>0 es) - - e<′₁e∷es : ∀ e (es : All (Expression Γ) Δ) → (t , e) <′₁ (e ∷ es) - e<′₁e∷es e es with callDepth e in eq₁ | callDepth′ es in eq₂ | ℕ.compare (callDepth e) (callDepth′ es) - ... | _ | _ | ℕ.less m k = inj₁ - (begin - suc (callDepth e) ≡⟨ cong suc eq₁ ⟩ - suc m ≤⟨ ℕₚ.m≤m+n (suc m) k ⟩ - suc m ℕ.+ k ≤⟨ ℕₚ.m≤n⊔m m (suc m ℕ.+ k) ⟩ - m ⊔ (suc m ℕ.+ k) ≡˘⟨ cong₂ _⊔_ eq₁ eq₂ ⟩ - callDepth e ⊔ callDepth′ es ∎) - ... | _ | _ | ℕ.equal m = inj₂ - (begin-equality - callDepth e ≡⟨ eq₁ ⟩ - m ≡˘⟨ ℕₚ.⊔-idem m ⟩ - m ⊔ m ≡˘⟨ cong₂ _⊔_ eq₁ eq₂ ⟩ - callDepth e ⊔ callDepth′ es ∎) - (e<ᵢe∷es e es) - ... | _ | _ | ℕ.greater n k = inj₂ - (sym (ℕₚ.m≥n⇒m⊔n≡m (begin - callDepth′ es ≡⟨ eq₂ ⟩ - n ≤⟨ ℕₚ.n≤1+n n ⟩ - suc n ≤⟨ ℕₚ.m≤m+n (suc n) k ⟩ - suc n ℕ.+ k ≡˘⟨ eq₁ ⟩ - callDepth e ∎))) - (e<ᵢe∷es e es) - - es<′₂e∷es : ∀ (e : Expression Γ t) (es : All (Expression Γ) Δ) → es <′₂ (e ∷ es) - es<′₂e∷es e es with callDepth e in eq₁ | callDepth′ es in eq₂ | ℕ.compare (callDepth e) (callDepth′ es) - ... | _ | _ | ℕ.less m k = inj₂ - (sym (ℕₚ.m≤n⇒m⊔n≡n (begin - callDepth e ≡⟨ eq₁ ⟩ - m ≤⟨ ℕₚ.n≤1+n m ⟩ - suc m ≤⟨ ℕₚ.m≤m+n (suc m) k ⟩ - suc m ℕ.+ k ≡˘⟨ eq₂ ⟩ - callDepth′ es ∎))) - (ℕₚ.m≤n+m (index′ es) (index e)) - ... | _ | _ | ℕ.equal m = inj₂ - (begin-equality - callDepth′ es ≡⟨ eq₂ ⟩ - m ≡˘⟨ ℕₚ.⊔-idem m ⟩ - m ⊔ m ≡˘⟨ cong₂ _⊔_ eq₁ eq₂ ⟩ - callDepth e ⊔ callDepth′ es ∎) - (ℕₚ.m≤n+m (index′ es) (index e)) - ... | _ | _ | ℕ.greater n k = inj₁ - (begin - suc (callDepth′ es) ≡⟨ cong suc eq₂ ⟩ - suc n ≤⟨ ℕₚ.m≤m+n (suc n) k ⟩ - suc n ℕ.+ k ≤⟨ ℕₚ.m≤m⊔n (suc n ℕ.+ k) n ⟩ - suc n ℕ.+ k ⊔ n ≡˘⟨ cong₂ _⊔_ eq₁ eq₂ ⟩ - callDepth e ⊔ callDepth′ es ∎) - - relWf = On.wellFounded < callDepth ∘ proj₂ , index ∘ proj₂ > (Lex.×-wellFounded ℕᵢ.<-wellFounded ℕᵢ.<-wellFounded) - - wf′₁ : ∀ f (es : All (Expression Γ) Δ) → Wf.WfRec rel pred (t , call f es) → - ∀ t,e → t,e <′₁ es → pred t,e - wf′₁ f _ rec (_ , e) r = rec (_ , e) (e<′₁es⇒e<f[es] r f) - - wf′₂ : ∀ f (es : All (Expression Γ) Δ) → Wf.WfRec rel pred (t , call f es) → - ∀ {n ts} (es′ : All (Expression Γ) {n} ts) → es′ <′₂ es → pred′ es′ - wf′₂ _ _ rec [] r = [] , refl - wf′₂ f es rec (e ∷ es′) r = proj₁ rec₁ ∷ proj₁ rec₂ , cong₂ _⊔_ (proj₂ rec₁) (proj₂ rec₂) - where - rec₁ = wf′₁ f es rec (_ , e) (<′₁-<′₂-trans (e<′₁e∷es e es′) r) - rec₂ = wf′₂ f es rec es′ (<′₂-trans (es<′₂e∷es e es′) r) - - rec₁ : ∀ (f : Expression Γ t → Expression Γ t′) → (∀ {e} → index (f e) ≡ suc (index e)) → (∀ {e} → callDepth (f e) ≡ callDepth e) → ∀ e → Wf.WfRec rel pred (t′ , f e) → pred (t′ , f e) - rec₁ f index-step depth-eq e acc = f (proj₁ inner) , trans depth-eq (proj₂ inner) - where inner = acc (_ , e) (inj₂ (sym depth-eq , ℕₚ.≤-reflexive (sym index-step))) - - rec₂ : ∀ {t′′} (f : Expression Γ t → Expression Γ t′ → Expression Γ t′′) → (∀ {e e₁} → index (f e e₁) ≡ suc (index e ℕ.+ index e₁)) → (∀ {e e₁} → callDepth (f e e₁) ≡ callDepth e ⊔ callDepth e₁) → ∀ e e₁ → Wf.WfRec rel pred (t′′ , f e e₁) → pred (t′′ , f e e₁) - rec₂ f index-step depth-eq e e₁ acc = f (proj₁ inner) (proj₁ inner₁) , trans depth-eq (cong₂ _⊔_ (proj₂ inner) (proj₂ inner₁)) - where - helper : ∀ a b c d e f → a ≡ b ⊔ c → d ≡ suc (e ℕ.+ f) → (b ℕ.< a ⊎ b ≡ a × e ℕ.< d) - helper a b c d e f eq₁ eq₂ with ℕ.compare b a - ... | ℕ.less .b k = inj₁ (ℕₚ.m≤m+n (suc b) k) - ... | ℕ.equal .a = inj₂ (refl , ℕₚ.m+n≤o⇒m≤o (suc e) (ℕₚ.≤-reflexive (sym eq₂))) - ... | ℕ.greater .a k = contradiction (ℕₚ.≤-reflexive (sym eq₁)) (ℕₚ.<⇒≱ (begin-strict - a <⟨ ℕₚ.n<1+n a ⟩ - suc a ≤⟨ ℕₚ.m≤m+n (suc a) k ⟩ - suc a ℕ.+ k ≤⟨ ℕₚ.m≤m⊔n (suc a ℕ.+ k) c ⟩ - suc a ℕ.+ k ⊔ c ∎)) - - helper₁ : ∀ a b c d e f → a ≡ b ⊔ c → d ≡ suc (e ℕ.+ f) → (c ℕ.< a ⊎ c ≡ a × f ℕ.< d) - helper₁ a b c d e f eq₁ eq₂ = helper a c b d f e (trans eq₁ (ℕₚ.⊔-comm b c)) (trans eq₂ (cong suc (ℕₚ.+-comm e f))) - - inner = acc (_ , e) (helper (callDepth (f e e₁)) (callDepth e) (callDepth e₁) (index (f e e₁)) (index e) (index e₁) depth-eq index-step) - inner₁ = acc (_ , e₁) (helper₁ (callDepth (f e e₁)) (callDepth e) (callDepth e₁) (index (f e e₁)) (index e) (index e₁) depth-eq index-step) - - rec₃ : ∀ {t′′ t′′′} (f : Expression Γ t → Expression Γ t′ → Expression Γ t′′ → Expression Γ t′′′) → (∀ {e e₁ e₂} → index (f e e₁ e₂) ≡ suc (index e ℕ.+ index e₁ ℕ.+ index e₂)) → (∀ {e e₁ e₂} → callDepth (f e e₁ e₂) ≡ callDepth e ⊔ callDepth e₁ ⊔ callDepth e₂) → ∀ e e₁ e₂ → Wf.WfRec rel pred (t′′′ , f e e₁ e₂) → pred (t′′′ , f e e₁ e₂) - rec₃ f index-step depth-eq e e₁ e₂ acc = f (proj₁ inner) (proj₁ inner₁) (proj₁ inner₂) , trans depth-eq (congₙ 3 (λ a b c → a ⊔ b ⊔ c) (proj₂ inner) (proj₂ inner₁) (proj₂ inner₂)) - where - helper : ∀ {a b c d e f g h} → a ≡ b ⊔ c ⊔ d → e ≡ suc (f ℕ.+ g ℕ.+ h) → b ℕ.< a ⊎ b ≡ a × f ℕ.< e - helper {a} {b} {c} {d} {e} {f} {g} {h} eq₁ eq₂ with ℕ.compare a b - ... | ℕ.less .a k = contradiction (ℕₚ.≤-reflexive (sym eq₁)) (ℕₚ.<⇒≱ (begin-strict - a <⟨ ℕₚ.n<1+n a ⟩ - suc a ≤⟨ ℕₚ.m≤m+n (suc a) k ⟩ - suc a ℕ.+ k ≤⟨ ℕₚ.m≤m⊔n (suc a ℕ.+ k) c ⟩ - suc a ℕ.+ k ⊔ c ≤⟨ ℕₚ.m≤m⊔n (suc a ℕ.+ k ⊔ c) d ⟩ - suc a ℕ.+ k ⊔ c ⊔ d ∎)) - ... | ℕ.equal .a = inj₂ (refl , (begin - suc f ≤⟨ ℕₚ.m≤m+n (suc f) g ⟩ - suc f ℕ.+ g ≤⟨ ℕₚ.m≤m+n (suc f ℕ.+ g) h ⟩ - suc f ℕ.+ g ℕ.+ h ≡˘⟨ eq₂ ⟩ - e ∎)) - ... | ℕ.greater .b k = inj₁ (ℕₚ.m≤m+n (suc b) k) - - helper₁ : ∀ {a} b {c} d {e} f {g} h → a ≡ b ⊔ c ⊔ d → e ≡ suc (f ℕ.+ g ℕ.+ h) → c ℕ.< a ⊎ c ≡ a × g ℕ.< e - helper₁ {a} b {c} d {e} f {g} h eq₁ eq₂ = - helper - (trans eq₁ (cong (_⊔ d) (ℕₚ.⊔-comm b c))) - (trans eq₂ (cong (ℕ._+ h) (cong suc (ℕₚ.+-comm f g)))) - - helper₂ : ∀ {a} b c {d e} f g {h} → a ≡ b ⊔ c ⊔ d → e ≡ suc (f ℕ.+ g ℕ.+ h) → d ℕ.< a ⊎ d ≡ a × h ℕ.< e - helper₂ {a} b c {d} {e} f g {h} eq₁ eq₂ = - helper - (trans eq₁ (trans (ℕₚ.⊔-comm (b ⊔ c) d) (sym (ℕₚ.⊔-assoc d b c)))) - (trans eq₂ (cong suc (trans (ℕₚ.+-comm (f ℕ.+ g) h) (sym (ℕₚ.+-assoc h f g))))) - - inner = acc (_ , e) (helper depth-eq index-step) - inner₁ = acc (_ , e₁) (helper₁ (callDepth e) (callDepth e₂) (index e) (index e₂) depth-eq index-step) - inner₂ = acc (_ , e₂) (helper₂ (callDepth e) (callDepth e₁) (index e) (index e₁) depth-eq index-step) - - helper : ∀ t,e → Wf.WfRec rel pred t,e → pred t,e - helper′ : ∀ (es : All (Expression Γ) Δ) → (∀ (t,e : ∃ (Expression Γ)) → t,e <′₁ es → pred t,e) → (∀ {n ts} (es′ : All (Expression Γ) {n} ts) → es′ <′₂ es → pred′ es′) → pred′ es - - helper (_ , Code.lit x) acc = lit x , refl - helper (_ , Code.state i) acc = state i , refl - helper (_ , Code.var i) acc = var i , refl - helper (_ , Code.abort e) acc = rec₁ abort refl refl e acc - helper (_ , _≟_ {hasEquality = hasEq} e e₁) acc = rec₂ (_≟_ {hasEquality = hasEq}) refl refl e e₁ acc - helper (_ , _<?_ {isNumeric = isNum} e e₁) acc = rec₂ (_<?_ {isNumeric = isNum}) refl refl e e₁ acc - helper (_ , Code.inv e) acc = rec₁ inv refl refl e acc - helper (_ , e Code.&& e₁) acc = rec₂ _&&_ refl refl e e₁ acc - helper (_ , e Code.|| e₁) acc = rec₂ _||_ refl refl e e₁ acc - helper (_ , Code.not e) acc = rec₁ not refl refl e acc - helper (_ , e Code.and e₁) acc = rec₂ _and_ refl refl e e₁ acc - helper (_ , e Code.or e₁) acc = rec₂ _or_ refl refl e e₁ acc - helper (_ , Code.[ e ]) acc = rec₁ [_] refl refl e acc - helper (_ , Code.unbox e) acc = rec₁ unbox refl refl e acc - helper (_ , Code.splice e e₁ e₂) acc = rec₃ splice refl refl e e₁ e₂ acc - helper (_ , Code.cut e e₁) acc = rec₂ cut refl refl e e₁ acc - helper (_ , Code.cast eq e) acc = rec₁ (cast eq) refl refl e acc - helper (_ , -_ {isNumeric = isNum} e) acc = rec₁ (-_ {isNumeric = isNum}) refl refl e acc - helper (_ , e Code.+ e₁) acc = rec₂ _+_ refl refl e e₁ acc - helper (_ , e Code.* e₁) acc = rec₂ _*_ refl refl e e₁ acc - helper (_ , _^_ {isNumeric = isNum} e x) acc = rec₁ (λ e → _^_ {isNumeric = isNum} e x) refl refl e acc - helper (_ , e Code.>> x) acc = rec₁ (_>> x) refl refl e acc - helper (_ , Code.rnd e) acc = rec₁ rnd refl refl e acc - helper (_ , Code.fin x e) acc = rec₁ (fin x) refl refl e acc - helper (_ , Code.asInt e) acc = rec₁ asInt refl refl e acc - helper (_ , Code.nil) acc = nil , refl - helper (_ , Code.cons e e₁) acc = rec₂ cons refl refl e e₁ acc - helper (_ , Code.head e) acc = rec₁ head refl refl e acc - helper (_ , Code.tail e) acc = rec₁ tail refl refl e acc - helper (_ , Code.call f es) acc = - acc - (_ , inlineFunction f (proj₁ inner)) +module Flatten where + private + structure : Expression Σ Γ t → ℕ + structures : All (Expression Σ Γ) ts → ℕ + structure (lit x) = suc (Σ[ 0 ] _) + structure (state i) = suc (Σ[ 0 ] _) + structure (var i) = suc (Σ[ 0 ] _) + structure (e ≟ e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (e <? e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (inv e) = suc (Σ[ 1 ] structure e) + structure (e && e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (e || e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (not e) = suc (Σ[ 1 ] structure e) + structure (e and e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (e or e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure ([ e ]) = suc (Σ[ 1 ] structure e) + structure (unbox e) = suc (Σ[ 1 ] structure e) + structure (merge e e₁ e₂) = suc (Σ[ 3 ] (structure e , structure e₁ , structure e₂)) + structure (slice e e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (cut e e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (cast eq e) = suc (Σ[ 1 ] structure e) + structure (- e) = suc (Σ[ 1 ] structure e) + structure (e + e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (e * e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (e ^ x) = suc (Σ[ 1 ] structure e) + structure (e >> n) = suc (Σ[ 1 ] structure e) + structure (rnd e) = suc (Σ[ 1 ] structure e) + structure (fin f e) = suc (Σ[ 1 ] structure e) + structure (asInt e) = suc (Σ[ 1 ] structure e) + structure (nil) = suc (Σ[ 0 ] _) + structure (cons e e₁) = suc (Σ[ 2 ] (structure e , structure e₁)) + structure (head e) = suc (Σ[ 1 ] structure e) + structure (tail e) = suc (Σ[ 1 ] structure e) + structure (call f es) = structures es + structure (if e then e₁ else e₂) = suc (Σ[ 3 ] (structure e , structure e₁ , structure e₂)) + + structures [] = 1 + structures (e ∷ es) = structures es ℕ.+ structure e + + structures>0 : ∀ (es : All (Expression Σ Γ) ts) → 0 < structures es + structures>0 [] = ℕₚ.0<1+n + structures>0 (e ∷ es) = ℕₚ.<-transˡ (structures>0 es) (ℕₚ.m≤m+n _ _) + + structure-∷ˡ-< : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → structure e < structures (e ∷ es) + structure-∷ˡ-< e es = ℕₚ.m<n+m _ (structures>0 es) + + structure-∷ʳ-≤ : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → structures es ≤ structures (e ∷ es) + structure-∷ʳ-≤ e es = ℕₚ.m≤m+n _ _ + + RecItem : Vec Type m → Vec Type n → Set + RecItem Σ Γ = ∃ (Expression Σ Γ) + + RecItems : Vec Type m → Vec Type n → Set + RecItems Σ Γ = ∃ λ n → ∃ (All (Expression Σ Γ) {n}) + + P : ∀ (Σ : Vec Type m) (Γ : Vec Type n) → RecItem Σ Γ → Set + P Σ Γ (t , e) = ∃ λ (e′ : Expression Σ Γ t) → CallDepth.expr e′ ≡ 0 + + Ps : ∀ (Σ : Vec Type k) (Γ : Vec Type m) → RecItems Σ Γ → Set + Ps Σ Γ (n , ts , es) = ∃ λ (es′ : All (Expression Σ Γ) ts) → CallDepth.exprs es′ ≡ 0 + + index : RecItem Σ Γ → ℕ × ℕ + index = < CallDepth.expr , structure > ∘ proj₂ + index′ : RecItems Σ Γ → ℕ × ℕ + index′ = < CallDepth.exprs , structures > ∘ proj₂ ∘ proj₂ + + infix 4 _≺_ _≺′_ _≺′′_ + + _≺_ : RecItem Σ Γ → RecItem Σ Γ → Set + _≺_ = ×-Lex _≡_ _<_ _<_ on index + + _≺′_ : RecItem Σ Γ → RecItems Σ Γ → Set + item ≺′ items = ×-Lex _≡_ _<_ _<_ (index item) (index′ items) + + _≺′′_ : RecItems Σ Γ → RecItems Σ Γ → Set + _≺′′_ = ×-Lex _≡_ _<_ _≤_ on index′ + + ≺-wellFounded : WellFounded (_≺_ {Σ = Σ} {Γ = Γ}) + ≺-wellFounded = On.wellFounded index (×-wellFounded ℕᵢ.<-wellFounded ℕᵢ.<-wellFounded) + + ≤∧<⇒≺ : ∀ (item item₁ : RecItem Σ Γ) → (_≤_ on proj₁ ∘ index) item item₁ → (_<_ on proj₂ ∘ index) item item₁ → item ≺ item₁ + ≤∧<⇒≺ item item₁ ≤₁ <₂ with proj₁ (index item) ℕₚ.<? proj₁ (index item₁) + ... | yes <₁ = inj₁ <₁ + ... | no ≮₁ = inj₂ (ℕₚ.≤∧≮⇒≡ ≤₁ ≮₁ , <₂) + + ≤∧<⇒≺′ : ∀ (item : RecItem Σ Γ) items → proj₁ (index item) ≤ proj₁ (index′ items) → proj₂ (index item) < proj₂ (index′ items) → item ≺′ items + ≤∧<⇒≺′ item items ≤₁ <₂ with proj₁ (index item) ℕₚ.<? proj₁ (index′ items) + ... | yes <₁ = inj₁ <₁ + ... | no ≮₁ = inj₂ (ℕₚ.≤∧≮⇒≡ ≤₁ ≮₁ , <₂) + + ≤∧≤⇒≺′′ : ∀ (items items₁ : RecItems Σ Γ) → (_≤_ on proj₁ ∘ index′) items items₁ → (_≤_ on proj₂ ∘ index′) items items₁ → items ≺′′ items₁ + ≤∧≤⇒≺′′ items items₁ ≤₁ ≤₂ with proj₁ (index′ items) ℕₚ.<? proj₁ (index′ items₁) + ... | yes <₁ = inj₁ <₁ + ... | no ≮₁ = inj₂ (ℕₚ.≤∧≮⇒≡ ≤₁ ≮₁ , ≤₂) + + ≺′-≺′′-trans : ∀ (item : RecItem Σ Γ) items items₁ → item ≺′ items → items ≺′′ items₁ → item ≺′ items₁ + ≺′-≺′′-trans _ _ _ (inj₁ <₁) (inj₁ <₂) = inj₁ (ℕₚ.<-trans <₁ <₂) + ≺′-≺′′-trans _ _ _ (inj₁ <₁) (inj₂ (≡₂ , _)) = inj₁ (proj₁ ℕₚ.<-resp₂-≡ ≡₂ <₁) + ≺′-≺′′-trans _ _ _ (inj₂ (≡₁ , _)) (inj₁ <₂) = inj₁ (proj₂ ℕₚ.<-resp₂-≡ (sym ≡₁) <₂) + ≺′-≺′′-trans _ _ _ (inj₂ (≡₁ , <₁)) (inj₂ (≡₂ , ≤₂)) = inj₂ (trans ≡₁ ≡₂ , ℕₚ.<-transˡ <₁ ≤₂) + + ≺′′-trans : ∀ (items items₁ items₂ : RecItems Σ Γ) → items ≺′′ items₁ → items₁ ≺′′ items₂ → items ≺′′ items₂ + ≺′′-trans _ _ _ (inj₁ <₁) (inj₁ <₂) = inj₁ (ℕₚ.<-trans <₁ <₂) + ≺′′-trans _ _ _ (inj₁ <₁) (inj₂ (≡₂ , _)) = inj₁ (proj₁ ℕₚ.<-resp₂-≡ ≡₂ <₁) + ≺′′-trans _ _ _ (inj₂ (≡₁ , _)) (inj₁ <₂) = inj₁ (proj₂ ℕₚ.<-resp₂-≡ (sym ≡₁) <₂) + ≺′′-trans _ _ _ (inj₂ (≡₁ , ≤₁)) (inj₂ (≡₂ , ≤₂)) = inj₂ (trans ≡₁ ≡₂ , ℕₚ.≤-trans ≤₁ ≤₂) + + ∷ˡ-≺′ : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → (-, e) ≺′ (-, -, e ∷ es) + ∷ˡ-≺′ e es = ≤∧<⇒≺′ (-, e) (-, -, e ∷ es) (CallDepth.∷ˡ-≤ e es) (structure-∷ˡ-< e es) + + ∷ʳ-≺′′ : ∀ (e : Expression Σ Γ t) (es : All (Expression Σ Γ) ts) → (-, -, es) ≺′′ (-, -, e ∷ es) + ∷ʳ-≺′′ e es = ≤∧≤⇒≺′′ (-, -, es) (-, -, e ∷ es) (CallDepth.∷ʳ-≤ e es) (structure-∷ʳ-≤ e es) + + toVecᵣ : ∀ {ts : Vec Type n} → All (Expression Σ Γ) ts → RecItem Σ Γ Vecᵣ.^ n + toVecᵣ [] = _ + toVecᵣ (e ∷ es) = Vecᵣ.cons _ (-, e) (toVecᵣ es) + + toSets : Vec Type m → Vec Type n → Vec Type o → Sets o 0ℓs + toSets Σ Γ [] = _ + toSets Σ Γ (t ∷ ts) = Expression Σ Γ t , toSets Σ Γ ts + + toProduct : ∀ {ts : Vec Type n} → All (Expression Σ Γ) ts → Product n (toSets Σ Γ ts) + toProduct [] = _ + toProduct (e ∷ []) = e + toProduct (e ∷ e₁ ∷ es) = e , toProduct (e₁ ∷ es) + + rec-helper : ∀ {Σ : Vec Type k} {Γ : Vec Type m} {ts : Vec Type n} + i (es : All (Expression Σ Γ) ts) (f : toSets Σ Γ ts ⇉ Expression Σ Γ t) → + (⨆[ n ] Vecᵣ.map (proj₁ ∘ index) n (toVecᵣ es) , suc (Σ[ n ] Vecᵣ.map (proj₂ ∘ index) n (toVecᵣ es))) ≡ index (-, uncurryₙ n f (toProduct es)) → + Vecᵣ.lookup i (toVecᵣ es) ≺ (-, uncurryₙ n f (toProduct es)) + rec-helper {n = n} i es f eq = ≤∧<⇒≺ + (Vecᵣ.lookup i (toVecᵣ es)) + (-, uncurryₙ n f (toProduct es)) + (begin + proj₁ (index (Vecᵣ.lookup i (toVecᵣ es))) ≡˘⟨ lookupᵣ-map i (toVecᵣ es) ⟩ + Vecᵣ.lookup i (Vecᵣ.map (proj₁ ∘ index) _ (toVecᵣ es)) ≤⟨ lookup-⨆-≤ i (Vecᵣ.map (proj₁ ∘ index) _ (toVecᵣ es)) ⟩ + ⨆[ n ] Vecᵣ.map (proj₁ ∘ index) _ (toVecᵣ es) ≡⟨ cong proj₁ eq ⟩ + CallDepth.expr (uncurryₙ n f (toProduct es)) ∎) + (begin-strict + proj₂ (index (Vecᵣ.lookup i (toVecᵣ es))) ≡˘⟨ lookupᵣ-map i (toVecᵣ es) ⟩ + Vecᵣ.lookup i (Vecᵣ.map (proj₂ ∘ index) _ (toVecᵣ es)) ≤⟨ lookup-Σ-≤ i (Vecᵣ.map (proj₂ ∘ index) _ (toVecᵣ es)) ⟩ + Σ[ n ] Vecᵣ.map (proj₂ ∘ index) _ (toVecᵣ es) <⟨ ℕₚ.n<1+n _ ⟩ + suc (Σ[ n ] Vecᵣ.map (proj₂ ∘ index) _ (toVecᵣ es)) ≡⟨ cong proj₂ eq ⟩ + structure (uncurryₙ n f (toProduct es)) ∎) + + helper : ∀ item → Wf.WfRec _≺_ (P Σ Γ) item → P Σ Γ item + helper′ : ∀ items → (∀ item → item ≺′ items → P Σ Γ item) → (∀ items₁ → items₁ ≺′′ items → Ps Σ Γ items₁) → Ps Σ Γ items + helper (_ , lit x) rec = lit x , refl + helper (_ , state i) rec = state i , refl + helper (_ , var i) rec = var i , refl + helper (_ , (e ≟ e₁)) rec = (proj₁ e′ ≟ proj₁ e₁′) , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _≟_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _≟_ refl) + helper (_ , (e <? e₁)) rec = (proj₁ e′ <? proj₁ e₁′) , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _<?_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _<?_ refl) + helper (_ , inv e) rec = inv (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) inv refl) + helper (_ , e && e₁) rec = proj₁ e′ && proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _&&_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _&&_ refl) + helper (_ , e || e₁) rec = proj₁ e′ || proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _||_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _||_ refl) + helper (_ , not e) rec = not (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) not refl) + helper (_ , e and e₁) rec = proj₁ e′ and proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _and_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _and_ refl) + helper (_ , e or e₁) rec = proj₁ e′ or proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _or_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _or_ refl) + helper (_ , [ e ]) rec = [ proj₁ e′ ] , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) [_] refl) + helper (_ , unbox e) rec = unbox (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) unbox refl) + helper (_ , merge e e₁ e₂) rec = merge (proj₁ e′) (proj₁ e₁′) (proj₁ e₂′) , congₙ 3 (λ a b c → a ⊔ b ⊔ c) (proj₂ e′) (proj₂ e₁′) (proj₂ e₂′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ e₂ ∷ []) merge refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ e₂ ∷ []) merge refl) + e₂′ = rec (-, e₂) (rec-helper 2F (e ∷ e₁ ∷ e₂ ∷ []) merge refl) + helper (_ , slice e e₁) rec = slice (proj₁ e′) (proj₁ e₁′) , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) slice refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) slice refl) + helper (_ , cut e e₁) rec = cut (proj₁ e′) (proj₁ e₁′) , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) cut refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) cut refl) + helper (_ , cast eq e) rec = cast eq (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) (cast eq) refl) + helper (_ , - e) rec = - proj₁ e′ , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) -_ refl) + helper (_ , e + e₁) rec = proj₁ e′ + proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _+_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _+_ refl) + helper (_ , e * e₁) rec = proj₁ e′ * proj₁ e₁′ , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) _*_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) _*_ refl) + helper (_ , e ^ x) rec = proj₁ e′ ^ x , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) (_^ x) refl) + helper (_ , e >> n) rec = proj₁ e′ >> n , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) (_>> n) refl) + helper (_ , rnd e) rec = rnd (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) rnd refl) + helper (_ , fin f e) rec = fin f (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) (fin f) refl) + helper (_ , asInt e) rec = asInt (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) asInt refl) + helper (_ , nil) rec = nil , refl + helper (_ , cons e e₁) rec = cons (proj₁ e′) (proj₁ e₁′) , cong₂ _⊔_ (proj₂ e′) (proj₂ e₁′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ []) cons refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ []) cons refl) + helper (_ , head e) rec = head (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) head refl) + helper (_ , tail e) rec = tail (proj₁ e′) , proj₂ e′ + where e′ = rec (-, e) (rec-helper 0F (e ∷ []) tail refl) + helper (_ , call f es) rec = rec + (-, Inline.fun f (proj₁ es′)) (inj₁ (begin-strict - callDepth (inlineFunction f (proj₁ inner)) ≤⟨ inlineFunction-lowers-callDepth f (proj₁ inner) ⟩ - funCallDepth f ⊔ callDepth′ (proj₁ inner) ≡⟨ cong (funCallDepth f ⊔_) (proj₂ inner) ⟩ - funCallDepth f ⊔ 0 ≡⟨ ℕₚ.⊔-identityʳ (funCallDepth f) ⟩ - funCallDepth f <⟨ ℕₚ.n<1+n (funCallDepth f) ⟩ - suc (funCallDepth f) ≤⟨ ℕₚ.m≤m⊔n (suc (funCallDepth f)) (callDepth′ es) ⟩ - suc (funCallDepth f) ⊔ callDepth′ es ∎)) - where inner = helper′ es (wf′₁ f es acc) (wf′₂ f es acc) - helper (_ , (Code.if e then e₁ else e₂)) acc = rec₃ if_then_else_ refl refl e e₁ e₂ acc - - proj₁ (helper′ [] acc acc′) = [] - proj₁ (helper′ (e ∷ es) acc acc′) = proj₁ (acc (_ , e) (e<′₁e∷es e es)) ∷ proj₁ (acc′ es (es<′₂e∷es e es)) - proj₂ (helper′ [] acc acc′) = refl - proj₂ (helper′ (e ∷ es) acc acc′) = cong₂ _⊔_ (proj₂ (acc (_ , e) (e<′₁e∷es e es))) (proj₂ (acc′ es (es<′₂e∷es e es))) + CallDepth.expr (Inline.fun f (proj₁ es′)) ≤⟨ Inline.fun-depth f (proj₁ es′) ⟩ + CallDepth.fun f ⊔ CallDepth.exprs (proj₁ es′) ≡⟨ cong (CallDepth.fun f ⊔_) (proj₂ es′) ⟩ + CallDepth.fun f ⊔ 0 ≡⟨ ℕₚ.⊔-identityʳ _ ⟩ + CallDepth.fun f <⟨ ℕₚ.n<1+n _ ⟩ + suc (CallDepth.fun f) ≤⟨ ℕₚ.m≤n⊔m (CallDepth.exprs es) _ ⟩ + CallDepth.exprs es ⊔ suc (CallDepth.fun f) ∎)) + where + rec′ : ∀ item → item ≺′ (-, (-, es)) → P _ _ item + rec′ item i≺es = rec item (lemma item i≺es) + where + lemma : ∀ item → item ≺′ (-, -, es) → item ≺ (-, call f es) + lemma item (inj₁ <-depth) = inj₁ (begin-strict + CallDepth.expr (proj₂ item) <⟨ <-depth ⟩ + CallDepth.exprs es ≤⟨ ℕₚ.m≤m⊔n (CallDepth.exprs es) _ ⟩ + CallDepth.expr (call f es) ∎) + lemma item (inj₂ (≡-depth , <-structure)) = ≤∧<⇒≺ item (-, call f es) + (begin + CallDepth.expr (proj₂ item) ≡⟨ ≡-depth ⟩ + CallDepth.exprs es ≤⟨ ℕₚ.m≤m⊔n (CallDepth.exprs es) _ ⟩ + CallDepth.expr (call f es) ∎) + <-structure + + rec′′ : ∀ items → items ≺′′ (-, (-, es)) → Ps _ _ items + rec′′ (_ , _ , es′) = go es′ + where + go : ∀ (es′ : All (Expression _ _) ts) → (-, -, es′) ≺′′ (-, -, es) → Ps _ _ (-, -, es′) + go [] ≺′′ = [] , refl + go (e ∷ es′) ≺′′ = proj₁ a ∷ proj₁ b , cong₂ _⊔_ (proj₂ b) (proj₂ a) + where + a = rec′ (-, e) (≺′-≺′′-trans (-, e) (-, -, e ∷ es′) (-, -, es) (∷ˡ-≺′ e es′) ≺′′) + b = go es′ (≺′′-trans (-, -, es′) (-, -, e ∷ es′) (-, -, es) (∷ʳ-≺′′ e es′) ≺′′) + + es′ = helper′ (-, -, es) rec′ rec′′ + helper (_ , (if e then e₁ else e₂)) rec = (if proj₁ e′ then proj₁ e₁′ else proj₁ e₂′) , congₙ 3 (λ a b c → a ⊔ b ⊔ c) (proj₂ e′) (proj₂ e₁′) (proj₂ e₂′) + where + e′ = rec (-, e) (rec-helper 0F (e ∷ e₁ ∷ e₂ ∷ []) if_then_else_ refl) + e₁′ = rec (-, e₁) (rec-helper 1F (e ∷ e₁ ∷ e₂ ∷ []) if_then_else_ refl) + e₂′ = rec (-, e₂) (rec-helper 2F (e ∷ e₁ ∷ e₂ ∷ []) if_then_else_ refl) + + helper′ (_ , _ , []) rec′ rec′′ = [] , refl + helper′ (_ , _ , e ∷ es) rec′ rec′′ = + proj₁ a ∷ proj₁ b , cong₂ _⊔_ (proj₂ b) (proj₂ a) + where + a = rec′ (-, e) (∷ˡ-≺′ e es) + b = rec′′ (-, -, es) (∷ʳ-≺′′ e es) + + expr : Expression Σ Γ t → Expression Σ Γ t + expr e = proj₁ (Wf.All.wfRec ≺-wellFounded _ (P _ _) helper (-, e)) + + expr-depth : ∀ (e : Expression Σ Γ t) → CallDepth.expr (expr e) ≡ 0 + expr-depth e = proj₂ (Wf.All.wfRec ≺-wellFounded _ (P _ _) helper (-, e)) diff --git a/src/Helium/Data/Pseudocode/Properties.agda b/src/Helium/Data/Pseudocode/Properties.agda deleted file mode 100644 index d73b4dd..0000000 --- a/src/Helium/Data/Pseudocode/Properties.agda +++ /dev/null @@ -1,109 +0,0 @@ ------------------------------------------------------------------------- --- Agda Helium --- --- Basic properties of the pseudocode data types ------------------------------------------------------------------------- - -{-# OPTIONS --without-K --safe #-} - -module Helium.Data.Pseudocode.Properties where - -import Data.Nat as ℕ -open import Data.Product using (_,_; uncurry) -open import Data.Vec using ([]; _∷_) -open import Function using (_∋_) -open import Helium.Data.Pseudocode.Core -import Relation.Binary.Consequences as Consequences -open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; cong₂) -open import Relation.Nullary using (Dec; yes; no) -open import Relation.Nullary.Decidable.Core using (map′) -open import Relation.Nullary.Product using (_×-dec_) - -infixl 4 _≟ᵗ_ _≟ˢ_ - -_≟ᵗ_ : ∀ (t t′ : Type) → Dec (t ≡ t′) -bool ≟ᵗ bool = yes refl -bool ≟ᵗ int = no (λ ()) -bool ≟ᵗ fin n = no (λ ()) -bool ≟ᵗ real = no (λ ()) -bool ≟ᵗ bit = no (λ ()) -bool ≟ᵗ bits n = no (λ ()) -bool ≟ᵗ tuple n x = no (λ ()) -bool ≟ᵗ array t′ n = no (λ ()) -int ≟ᵗ bool = no (λ ()) -int ≟ᵗ int = yes refl -int ≟ᵗ fin n = no (λ ()) -int ≟ᵗ real = no (λ ()) -int ≟ᵗ bit = no (λ ()) -int ≟ᵗ bits n = no (λ ()) -int ≟ᵗ tuple n x = no (λ ()) -int ≟ᵗ array t′ n = no (λ ()) -fin n ≟ᵗ bool = no (λ ()) -fin n ≟ᵗ int = no (λ ()) -fin m ≟ᵗ fin n = map′ (cong fin) (λ { refl → refl }) (m ℕ.≟ n) -fin n ≟ᵗ real = no (λ ()) -fin n ≟ᵗ bit = no (λ ()) -fin n ≟ᵗ bits n₁ = no (λ ()) -fin n ≟ᵗ tuple n₁ x = no (λ ()) -fin n ≟ᵗ array t′ n₁ = no (λ ()) -real ≟ᵗ bool = no (λ ()) -real ≟ᵗ int = no (λ ()) -real ≟ᵗ fin n = no (λ ()) -real ≟ᵗ real = yes refl -real ≟ᵗ bit = no (λ ()) -real ≟ᵗ bits n = no (λ ()) -real ≟ᵗ tuple n x = no (λ ()) -real ≟ᵗ array t′ n = no (λ ()) -bit ≟ᵗ bool = no (λ ()) -bit ≟ᵗ int = no (λ ()) -bit ≟ᵗ fin n = no (λ ()) -bit ≟ᵗ real = no (λ ()) -bit ≟ᵗ bit = yes refl -bit ≟ᵗ bits n = no (λ ()) -bit ≟ᵗ tuple n x = no (λ ()) -bit ≟ᵗ array t n = no (λ ()) -bits n ≟ᵗ bool = no (λ ()) -bits n ≟ᵗ int = no (λ ()) -bits n ≟ᵗ fin n₁ = no (λ ()) -bits n ≟ᵗ real = no (λ ()) -bits m ≟ᵗ bit = no (λ ()) -bits m ≟ᵗ bits n = map′ (cong bits) (λ { refl → refl }) (m ℕ.≟ n) -bits n ≟ᵗ tuple n₁ x = no (λ ()) -bits n ≟ᵗ array t′ n₁ = no (λ ()) -tuple n x ≟ᵗ bool = no (λ ()) -tuple n x ≟ᵗ int = no (λ ()) -tuple n x ≟ᵗ fin n₁ = no (λ ()) -tuple n x ≟ᵗ real = no (λ ()) -tuple n x ≟ᵗ bit = no (λ ()) -tuple n x ≟ᵗ bits n₁ = no (λ ()) -tuple _ [] ≟ᵗ tuple _ [] = yes refl -tuple _ [] ≟ᵗ tuple _ (y ∷ ys) = no (λ ()) -tuple _ (x ∷ xs) ≟ᵗ tuple _ [] = no (λ ()) -tuple _ (x ∷ xs) ≟ᵗ tuple _ (y ∷ ys) = map′ (λ { (refl , refl) → refl }) (λ { refl → refl , refl }) (x ≟ᵗ y ×-dec tuple _ xs ≟ᵗ tuple _ ys) -tuple n x ≟ᵗ array t′ n₁ = no (λ ()) -array t n ≟ᵗ bool = no (λ ()) -array t n ≟ᵗ int = no (λ ()) -array t n ≟ᵗ fin n₁ = no (λ ()) -array t n ≟ᵗ real = no (λ ()) -array t n ≟ᵗ bit = no (λ ()) -array t n ≟ᵗ bits n₁ = no (λ ()) -array t n ≟ᵗ tuple n₁ x = no (λ ()) -array t m ≟ᵗ array t′ n = map′ (uncurry (cong₂ array)) (λ { refl → refl , refl }) (t ≟ᵗ t′ ×-dec m ℕ.≟ n) - -_≟ˢ_ : ∀ (t t′ : Sliced) → Dec (t ≡ t′) -bits ≟ˢ bits = yes refl -bits ≟ˢ array x = no (λ ()) -array x ≟ˢ bits = no (λ ()) -array x ≟ˢ array y = map′ (cong array) (λ { refl → refl }) (x ≟ᵗ y) - -bits-injective : ∀ {m n} → (Type ∋ bits m) ≡ bits n → m ≡ n -bits-injective refl = refl - -array-injective₁ : ∀ {t t′ m n} → (Type ∋ array t m) ≡ array t′ n → t ≡ t′ -array-injective₁ refl = refl - -array-injective₂ : ∀ {t t′ m n} → (Type ∋ array t m) ≡ array t′ n → m ≡ n -array-injective₂ refl = refl - -typeEqRecomp : ∀ {t t′} → .(eq : t ≡ t′) → t ≡ t′ -typeEqRecomp = Consequences.dec⇒recomputable _≟ᵗ_ diff --git a/src/Helium/Instructions/Base.agda b/src/Helium/Instructions/Base.agda index d157d5a..3e1bb5f 100644 --- a/src/Helium/Instructions/Base.agda +++ b/src/Helium/Instructions/Base.agda @@ -11,18 +11,25 @@ module Helium.Instructions.Base where open import Data.Bool as Bool using (true; false) open import Data.Fin as Fin using (Fin; Fin′; zero; suc; toℕ) open import Data.Fin.Patterns +open import Data.Integer as ℤ using (1ℤ; 0ℤ; -1ℤ) open import Data.Nat as ℕ using (ℕ; zero; suc) import Data.Nat.Properties as ℕₚ +open import Data.Product using (uncurry) open import Data.Sum using ([_,_]′; inj₂) open import Data.Vec as Vec using (Vec; []; _∷_) open import Data.Vec.Relation.Unary.All using (All; []; _∷_) open import Function using (_$_) open import Helium.Data.Pseudocode.Core as Core public - hiding (module Code) import Helium.Instructions.Core as Instr import Relation.Binary.PropositionalEquality as P open import Relation.Nullary.Decidable.Core using (True) +private + variable + k m n : ℕ + t : Type + Γ : Vec Type n + --- Types beat : Type @@ -43,75 +50,85 @@ State = array (bits 32) 32 -- S ∷ beat -- _BeatId ∷ [] -open Core.Code State public - --- References -- Direct from State -S : ∀ {n Γ} → Expression {n} Γ (array (bits 32) 32) +S : Reference State Γ (array (bits 32) 32) S = state 0F -R : ∀ {n Γ} → Expression {n} Γ (array (bits 32) 16) +R : Reference State Γ (array (bits 32) 16) R = state 1F -VPR-P0 : ∀ {n Γ} → Expression {n} Γ (bits 16) +VPR-P0 : Reference State Γ (bits 16) VPR-P0 = state 2F -VPR-mask : ∀ {n Γ} → Expression {n} Γ (bits 8) +VPR-mask : Reference State Γ (bits 8) VPR-mask = state 3F -FPSCR-QC : ∀ {n Γ} → Expression {n} Γ bit +FPSCR-QC : Reference State Γ bit FPSCR-QC = state 4F -AdvanceVPTState : ∀ {n Γ} → Expression {n} Γ bool +AdvanceVPTState : Reference State Γ bool AdvanceVPTState = state 5F -BeatId : ∀ {n Γ} → Expression {n} Γ beat +BeatId : Reference State Γ beat BeatId = state 6F -- Indirect -group : ∀ {n Γ t k} m → Expression {n} Γ (asType t (k ℕ.* suc m)) → Expression Γ (array (asType t k) (suc m)) -group {k = k} zero x = [ cast (P.trans (ℕₚ.*-comm k 1) (ℕₚ.+-comm k 0)) x ] -group {k = k} (suc m) x = group m (slice x′ (lit (Fin.fromℕ k ′f))) ∶ [ slice (cast (ℕₚ.+-comm k _) x′) (lit (zero ′f)) ] +index : Expression State Γ (array t (suc m)) → Expression State Γ (fin (suc m)) → Expression State Γ t +index {m = m} x i = unbox (slice (cast (ℕₚ.+-comm 1 m) x) i) + +*index : Reference State Γ (array t (suc m)) → Expression State Γ (fin (suc m)) → Reference State Γ t +*index {m = m} x i = unbox (slice (cast (ℕₚ.+-comm 1 m) x) i) + +*index-group : Reference State Γ (array t (k ℕ.* suc m)) → Expression State Γ (fin (suc m)) → Reference State Γ (array t k) +*index-group {k = k} {m = m} x i = slice (cast eq x) (fin reindex (tup (i ∷ []))) where - x′ = cast (P.trans (ℕₚ.*-comm k _) (P.cong (k ℕ.+_) (ℕₚ.*-comm _ k))) x + eq = P.trans (ℕₚ.*-comm k (suc m)) (ℕₚ.+-comm k (m ℕ.* k)) + + reindex : ∀ {m n} → Fin (suc m) → Fin (suc (m ℕ.* n)) + reindex {m} {n} 0F = Fin.inject+ (m ℕ.* n) 0F + reindex {suc m} {n} (suc i) = Fin.cast (ℕₚ.+-suc n (m ℕ.* n)) (Fin.raise n (reindex i)) -join : ∀ {n Γ t k m} → Expression {n} Γ (array (asType t k) (suc m)) → Expression Γ (asType t (k ℕ.* suc m)) -join {k = k} {zero} x = cast (P.trans (ℕₚ.+-comm 0 k) (ℕₚ.*-comm 1 k)) (unbox x) -join {k = k} {suc m} x = cast eq (join (slice x (lit (Fin.fromℕ 1 ′f))) ∶ unbox (slice {i = suc m} (cast (ℕₚ.+-comm 1 _) x) (lit (zero ′f)))) +index-group : Expression State Γ (array t (k ℕ.* suc m)) → Expression State Γ (fin (suc m)) → Expression State Γ (array t k) +index-group {k = k} {m = m} x i = slice (cast eq x) (fin reindex (tup (i ∷ []))) where - eq = P.trans (P.cong (k ℕ.+_) (ℕₚ.*-comm k (suc m))) (ℕₚ.*-comm (suc (suc m)) k) + eq = P.trans (ℕₚ.*-comm k (suc m)) (ℕₚ.+-comm k (m ℕ.* k)) -index : ∀ {n Γ t m} → Expression {n} Γ (asType t (suc m)) → Expression Γ (fin (suc m)) → Expression Γ (elemType t) -index {m = m} x i = unbox (slice (cast (ℕₚ.+-comm 1 m) x) i) + reindex : ∀ {m n} → Fin (suc m) → Fin (suc (m ℕ.* n)) + reindex {m} {n} 0F = Fin.inject+ (m ℕ.* n) 0F + reindex {suc m} {n} (suc i) = Fin.cast (ℕₚ.+-suc n (m ℕ.* n)) (Fin.raise n (reindex i)) + +Q[_,_] : Expression State Γ (fin 8) → Expression State Γ (fin 4) → Reference State Γ (bits 32) +Q[ i , j ] = *index S (fin (uncurry Fin.combine) (tup (i ∷ j ∷ []))) -Q : ∀ {n Γ} → Expression {n} Γ (array (array (bits 32) 4) 8) -Q = group 7 S +elem : ∀ m → Expression State Γ (array t (suc k ℕ.* m)) → Expression State Γ (fin (suc k)) → Expression State Γ (array t m) +elem {k = k} zero x i = cast (ℕₚ.*-comm k 0) x +elem {k = k} (suc m) x i = index-group (cast (ℕₚ.*-comm (suc k) (suc m)) x) i -elem : ∀ {n Γ t k} m → Expression {n} Γ (asType t (k ℕ.* m)) → Expression Γ (fin k) → Expression Γ (asType t m) -elem {k = zero} m x i = abort i -elem {k = suc k} zero x i = cast (ℕₚ.*-comm k 0) x -elem {k = suc k} (suc m) x i = index (group k (cast (ℕₚ.*-comm (suc k) (suc m)) x)) i +*elem : ∀ m → Reference State Γ (array t (suc k ℕ.* m)) → Expression State Γ (fin (suc k)) → Reference State Γ (array t m) +*elem {k = k} zero x i = cast (ℕₚ.*-comm k 0) x +*elem {k = k} (suc m) x i = *index-group (cast (ℕₚ.*-comm (suc k) (suc m)) x) i --- Other utiliies -hasBit : ∀ {n Γ m} → Expression {n} Γ (bits (suc m)) → Expression Γ (fin (suc m)) → Expression Γ bool -hasBit {n} x i = index x i ≟ lit (true ′x) +hasBit : Expression State Γ (bits (suc m)) → Expression State Γ (fin (suc m)) → Expression State Γ bool +hasBit {n} x i = index x i ≟ lit true -sliceⁱ : ∀ {n Γ m} → ℕ → Expression {n} Γ int → Expression Γ (bits m) -sliceⁱ {m = zero} n i = lit ([] ′xs) -sliceⁱ {m = suc m} n i = sliceⁱ (suc n) i ∶ [ get n i ] +sliceⁱ : ℕ → Expression State Γ int → Expression State Γ (bits m) +sliceⁱ {m = zero} n i = lit [] +sliceⁱ {m = suc m} n i = sliceⁱ (suc n) i ∶ [ getBit n i ] --- Functions -Int : ∀ {n} → Function (bits n ∷ bool ∷ []) int +Int : Function State (bits n ∷ bool ∷ []) int Int = skip ∙return (if var 1F then uint (var 0F) else sint (var 0F)) -- arguments swapped, pred n -SignedSatQ : ∀ n → Function (int ∷ []) (tuple 2 (bits (suc n) ∷ bool ∷ [])) -SignedSatQ n = declare (lit (true ′b)) ( +SignedSatQ : ∀ n → Function State (int ∷ []) (tuple (bits (suc n) ∷ bool ∷ [])) +SignedSatQ n = declare (lit true) ( if max <? var 1F then var 1F ≔ max @@ -119,18 +136,18 @@ SignedSatQ n = declare (lit (true ′b)) ( then var 1F ≔ min else - var 0F ≔ lit (false ′b) + var 0F ≔ lit false ∙return tup (sliceⁱ 0 (var 1F) ∷ var 0F ∷ [])) where - max = lit (2 ′i) ^ n + - lit (1 ′i) - min = - (lit (2 ′i) ^ n) + max = lit (ℤ.+ (2 ℕ.^ n) ℤ.+ -1ℤ) + min = lit (ℤ.- ℤ.+ (2 ℕ.^ n)) -- actual shift if 'shift + 1' -LSL-C : ∀ {n} (shift : ℕ) → Function (bits n ∷ []) (tuple 2 (bits n ∷ bit ∷ [])) -LSL-C {n} shift = declare (var 0F ∶ lit ((Vec.replicate {n = (suc shift)} false) ′xs)) +LSL-C : ∀ (shift : ℕ) → Function State (bits n ∷ []) (tuple (bits n ∷ bit ∷ [])) +LSL-C {n} shift = declare (var 0F ∶ lit ((Vec.replicate {n = (suc shift)} false))) (skip ∙return tup - ( slice (var 0F) (lit (zero ′f)) - ∷ unbox (slice (cast eq (var 0F)) (lit (Fin.inject+ shift (Fin.fromℕ n) ′f))) + ( slice (var 0F) (lit 0F) + ∷ unbox (slice (cast eq (var 0F)) (lit (Fin.inject+ shift (Fin.fromℕ n)))) ∷ [])) where eq = P.trans (ℕₚ.+-comm 1 (shift ℕ.+ n)) (P.cong (ℕ._+ 1) (ℕₚ.+-comm shift n)) @@ -138,49 +155,50 @@ LSL-C {n} shift = declare (var 0F ∶ lit ((Vec.replicate {n = (suc shift)} fals --- Procedures private - div2 : All Fin (4 ∷ []) → Fin 2 - div2 (zero ∷ []) = zero - div2 (suc zero ∷ []) = zero - div2 (suc (suc i) ∷ []) = suc zero + div2 : Fin 4 → Fin 2 + div2 0F = 0F + div2 1F = 0F + div2 2F = 1F + div2 3F = 1F -copyMasked : Procedure (fin 8 ∷ bits 32 ∷ beat ∷ elmtMask ∷ []) +copyMasked : Procedure State (fin 8 ∷ bits 32 ∷ beat ∷ elmtMask ∷ []) copyMasked = for 4 -- 0:e 1:dest 2:result 3:beat 4:elmtMask ( if hasBit (var 4F) (var 0F) then - elem 8 (index (index Q (var 1F)) (var 3F)) (var 0F) ≔ elem 8 (var 2F) (var 0F) + *elem 8 Q[ var 1F , var 3F ] (var 0F) ≔ elem 8 (var 2F) (var 0F) ) ∙end -VPTAdvance : Procedure (beat ∷ []) +VPTAdvance : Procedure State (beat ∷ []) VPTAdvance = declare (fin div2 (tup (var 0F ∷ []))) ( - declare (elem 4 VPR-mask (var 0F)) ( + declare (elem 4 (! VPR-mask) (var 0F)) ( -- 0:vptState 1:maskId 2:beat - if var 0F ≟ lit ((true ∷ false ∷ false ∷ false ∷ []) ′xs) + if var 0F ≟ lit (true ∷ false ∷ false ∷ false ∷ []) then - var 0F ≔ lit (Vec.replicate false ′xs) - else if inv (var 0F ≟ lit (Vec.replicate false ′xs)) + var 0F ≔ lit (Vec.replicate false) + else if inv (var 0F ≟ lit (Vec.replicate false)) then ( - declare (lit (false ′x)) ( + declare (lit false) ( -- 0:inv 1:vptState 2:maskId 3:beat - tup (var 1F ∷ var 0F ∷ []) ≔ call (LSL-C 0) (var 1F ∷ []) ∙ - if var 0F ≟ lit (true ′x) + cons (var 1F) (cons (var 0F) nil) ≔ call (LSL-C 0) (var 1F ∷ []) ∙ + if var 0F ≟ lit true then - elem 4 VPR-P0 (var 3F) ≔ not (elem 4 VPR-P0 (var 3F)))) ∙ - if get 0 (asInt (var 2F)) ≟ lit (true ′x) + *elem 4 VPR-P0 (var 3F) ≔ not (elem 4 (! VPR-P0) (var 3F)))) ∙ + if getBit 0 (asInt (var 2F)) ≟ lit true then - elem 4 VPR-mask (var 1F) ≔ var 0F)) + *elem 4 VPR-mask (var 1F) ≔ var 0F)) ∙end -VPTActive : Function (beat ∷ []) bool -VPTActive = skip ∙return inv (elem 4 VPR-mask (fin div2 (tup (var 0F ∷ []))) ≟ lit (Vec.replicate false ′xs)) +VPTActive : Function State (beat ∷ []) bool +VPTActive = skip ∙return inv (elem 4 (! VPR-mask) (fin div2 (tup (var 0F ∷ []))) ≟ lit (Vec.replicate false)) -GetCurInstrBeat : Function [] (tuple 2 (beat ∷ elmtMask ∷ [])) -GetCurInstrBeat = declare (lit (Vec.replicate true ′xs)) ( +GetCurInstrBeat : Function State [] (tuple (beat ∷ elmtMask ∷ [])) +GetCurInstrBeat = declare (lit (Vec.replicate true)) ( -- 0:elmtMask 1:beat - if call VPTActive (BeatId ∷ []) + if call VPTActive (! BeatId ∷ []) then - var 0F ≔ var 0F and elem 4 VPR-P0 BeatId - ∙return tup (BeatId ∷ var 0F ∷ [])) + var 0F ≔ var 0F and elem 4 (! VPR-P0) (! BeatId) + ∙return tup (! BeatId ∷ var 0F ∷ [])) -- Assumes: -- MAX_OVERLAPPING_INSTRS = 1 @@ -188,113 +206,113 @@ GetCurInstrBeat = declare (lit (Vec.replicate true ′xs)) ( -- BEATS_PER_TICK = 4 -- procedure argument is action of DecodeExecute -- and more! -ExecBeats : Procedure [] → Procedure [] +ExecBeats : Procedure State [] → Procedure State [] ExecBeats DecodeExec = for 4 ( -- 0:beatId BeatId ≔ var 0F ∙ - AdvanceVPTState ≔ lit (true ′b) ∙ + AdvanceVPTState ≔ lit true ∙ invoke DecodeExec [] ∙ - if AdvanceVPTState + if ! AdvanceVPTState then invoke VPTAdvance (var 0F ∷ [])) ∙end -from32 : ∀ size {n Γ} → Expression {n} Γ (bits 32) → Expression Γ (array (bits (toℕ (Instr.Size.esize size))) (toℕ (Instr.Size.elements size))) -from32 Instr.8bit = group 3 -from32 Instr.16bit = group 1 -from32 Instr.32bit = group 0 +*index-32 : ∀ size → Reference State Γ (bits 32) → Expression State Γ (fin (toℕ (Instr.Size.elements size))) → Reference State Γ (bits (toℕ (Instr.Size.esize size))) +*index-32 Instr.8bit = *index-group +*index-32 Instr.16bit = *index-group +*index-32 Instr.32bit = *index-group -to32 : ∀ size {n Γ} → Expression {n} Γ (array (bits (toℕ (Instr.Size.esize size))) (toℕ (Instr.Size.elements size))) → Expression Γ (bits 32) -to32 Instr.8bit = join -to32 Instr.16bit = join -to32 Instr.32bit = join +index-32 : ∀ size → Expression State Γ (bits 32) → Expression State Γ (fin (toℕ (Instr.Size.elements size))) → Expression State Γ (bits (toℕ (Instr.Size.esize size))) +index-32 Instr.8bit = index-group +index-32 Instr.16bit = index-group +index-32 Instr.32bit = index-group module _ (d : Instr.VecOp₂) where open Instr.VecOp₂ d -- 0:op₂ 1:e 2:op₁ 3:result 4:elmtMask 5:curBeat - vec-op₂′ : Statement (bits (toℕ esize) ∷ fin (toℕ elements) ∷ array (bits (toℕ esize)) (toℕ elements) ∷ array (bits (toℕ esize)) (toℕ elements) ∷ elmtMask ∷ beat ∷ []) → Procedure [] - vec-op₂′ op = declare (lit (zero ′f)) ( - declare (lit (Vec.replicate false ′xs)) ( + vec-op₂′ : Statement State (bits (toℕ esize) ∷ fin (toℕ elements) ∷ bits 32 ∷ bits 32 ∷ elmtMask ∷ beat ∷ []) → Procedure State [] + vec-op₂′ op = declare (lit 0F) ( + declare (lit (Vec.replicate false)) ( -- 0:elmtMask 1:curBeat - tup (var 1F ∷ var 0F ∷ []) ≔ call GetCurInstrBeat [] ∙ - declare (lit ((Vec.replicate false ′xs) ′a)) ( - declare (from32 size (index (index Q (lit (src₁ ′f))) (var 2F))) ( + cons (var 1F) (cons (var 0F) nil) ≔ call GetCurInstrBeat [] ∙ + declare (lit (Vec.replicate false)) ( + declare (! Q[ lit src₁ , var 2F ]) ( -- 0:op₁ 1:result 2:elmtMask 3:curBeat for (toℕ elements) ( -- 0:e 1:op₁ 2:result 3:elmtMask 4:curBeat declare op₂ op ) ∙ -- 0:op₁ 1:result 2:elmtMask 3:curBeat - invoke copyMasked (lit (dest ′f) ∷ to32 size (var 1F) ∷ var 3F ∷ var 2F ∷ []))))) + invoke copyMasked (lit dest ∷ var 1F ∷ var 3F ∷ var 2F ∷ []))))) ∙end where -- 0:e 1:op₁ 2:result 3:elmtMask 4:curBeat op₂ = - [ (λ src₂ → index (from32 size (index R (lit (src₂ ′f)))) (lit (zero ′f))) - , (λ src₂ → index (from32 size (index (index Q (lit (src₂ ′f))) (var 4F))) (var 0F)) + [ (λ src₂ → index-32 size (index (! R) (lit src₂)) (lit 0F)) + , (λ src₂ → index-32 size (! Q[ lit src₂ , var 4F ]) (var 0F)) ]′ src₂ - vec-op₂ : Function (bits (toℕ esize) ∷ bits (toℕ esize) ∷ []) (bits (toℕ esize)) → Procedure [] - vec-op₂ op = vec-op₂′ (index (var 3F) (var 1F) ≔ call op (index (var 2F) (var 1F) ∷ var 0F ∷ [])) + vec-op₂ : Function State (bits (toℕ esize) ∷ bits (toℕ esize) ∷ []) (bits (toℕ esize)) → Procedure State [] + vec-op₂ op = vec-op₂′ (*index-32 size (var 3F) (var 1F) ≔ call op (index-32 size (var 2F) (var 1F) ∷ var 0F ∷ [])) -vadd : Instr.VAdd → Procedure [] +vadd : Instr.VAdd → Procedure State [] vadd d = vec-op₂ d (skip ∙return sliceⁱ 0 (uint (var 0F) + uint (var 1F))) -vsub : Instr.VSub → Procedure [] +vsub : Instr.VSub → Procedure State [] vsub d = vec-op₂ d (skip ∙return sliceⁱ 0 (uint (var 0F) - uint (var 1F))) -vhsub : Instr.VHSub → Procedure [] +vhsub : Instr.VHSub → Procedure State [] vhsub d = vec-op₂ op₂ (skip ∙return sliceⁱ 1 (toInt (var 0F) - toInt (var 1F))) - where open Instr.VHSub d; toInt = λ i → call Int (i ∷ lit (unsigned ′b) ∷ []) + where open Instr.VHSub d; toInt = λ i → call Int (i ∷ lit unsigned ∷ []) -vmul : Instr.VMul → Procedure [] +vmul : Instr.VMul → Procedure State [] vmul d = vec-op₂ d (skip ∙return sliceⁱ 0 (sint (var 0F) * sint (var 1F))) -vmulh : Instr.VMulH → Procedure [] +vmulh : Instr.VMulH → Procedure State [] vmulh d = vec-op₂ op₂ (skip ∙return sliceⁱ (toℕ esize) (toInt (var 0F) * toInt (var 1F))) where - open Instr.VMulH d; toInt = λ i → call Int (i ∷ lit (unsigned ′b) ∷ []) + open Instr.VMulH d; toInt = λ i → call Int (i ∷ lit unsigned ∷ []) -vrmulh : Instr.VRMulH → Procedure [] -vrmulh d = vec-op₂ op₂ (skip ∙return sliceⁱ (toℕ esize) (toInt (var 0F) * toInt (var 1F) + lit (1 ′i) << toℕ esize-1)) +vrmulh : Instr.VRMulH → Procedure State [] +vrmulh d = vec-op₂ op₂ (skip ∙return sliceⁱ (toℕ esize) (toInt (var 0F) * toInt (var 1F) + lit 1ℤ << toℕ esize-1)) where - open Instr.VRMulH d; toInt = λ i → call Int (i ∷ lit (unsigned ′b) ∷ []) + open Instr.VRMulH d; toInt = λ i → call Int (i ∷ lit unsigned ∷ []) -vmla : Instr.VMlA → Procedure [] +vmla : Instr.VMlA → Procedure State [] vmla d = vec-op₂ op₂ (skip ∙return sliceⁱ (toℕ esize) (toInt (var 0F) * element₂ + toInt (var 1F))) where open Instr.VMlA d op₂ = record { size = size ; dest = acc ; src₁ = src₁ ; src₂ = inj₂ acc } - toInt = λ i → call Int (i ∷ lit (unsigned ′b) ∷ []) - element₂ = toInt (index (from32 size (index R (lit (src₂ ′f)))) (lit (zero ′f))) + toInt = λ i → call Int (i ∷ lit unsigned ∷ []) + element₂ = toInt (index-32 size (index (! R) (lit src₂)) (lit 0F)) private - vqr?dmulh : Instr.VQDMulH → Function (int ∷ int ∷ []) int → Procedure [] + vqr?dmulh : Instr.VQDMulH → Function State (int ∷ int ∷ []) int → Procedure State [] vqr?dmulh d f = vec-op₂′ d ( -- 0:op₂ 1:e 2:op₁ 3:result 4:elmtMask 5:curBeat - declare (call f (sint (index (var 2F) (var 1F)) ∷ sint (var 0F) ∷ [])) ( - declare (lit (false ′b)) ( + declare (call f (sint (index-32 size (var 2F) (var 1F)) ∷ sint (var 0F) ∷ [])) ( + declare (lit false) ( -- 0:sat 1:value 2:op₂ 3:e 4:op₁ 5:result 6:elmtMask 7:curBeat - tup (index (var 5F) (var 3F) ∷ var 0F ∷ []) ≔ call (SignedSatQ (toℕ esize-1)) (var 1F ∷ []) ∙ + cons (*index-32 size (var 5F) (var 3F)) (cons (var 0F) nil) ≔ call (SignedSatQ (toℕ esize-1)) (var 1F ∷ []) ∙ if var 0F && hasBit (var 6F) (fin e*esize>>3 (tup ((var 3F) ∷ []))) then - FPSCR-QC ≔ lit (true ′x)))) + FPSCR-QC ≔ lit true))) where open Instr.VecOp₂ d - e*esize>>3 : All Fin (toℕ elements ∷ []) → Fin 4 - e*esize>>3 (x ∷ []) = helper size x + e*esize>>3 : Fin (toℕ elements) → Fin 4 + e*esize>>3 x = helper size x where helper : ∀ size → Fin′ (Instr.Size.elements size) → Fin 4 helper Instr.8bit i = Fin.combine i (zero {0}) helper Instr.16bit i = Fin.combine i (zero {1}) helper Instr.32bit i = Fin.combine i zero -vqdmulh : Instr.VQDMulH → Procedure [] -vqdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0F * var 1F >> toℕ esize) +vqdmulh : Instr.VQDMulH → Procedure State [] +vqdmulh d = vqr?dmulh d (skip ∙return lit (ℤ.+ 2) * var 0F * var 1F >> toℕ esize) where open Instr.VecOp₂ d using (esize) -vqrdmulh : Instr.VQRDMulH → Procedure [] -vqrdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0F * var 1F + lit (1 ′i) << toℕ esize-1 >> toℕ esize) +vqrdmulh : Instr.VQRDMulH → Procedure State [] +vqrdmulh d = vqr?dmulh d (skip ∙return lit (ℤ.+ 2) * var 0F * var 1F + lit 1ℤ << toℕ esize-1 >> toℕ esize) where open Instr.VecOp₂ d using (esize; esize-1) diff --git a/src/Helium/Instructions/Instances/Barrett.agda b/src/Helium/Instructions/Instances/Barrett.agda index 606a9e9..b913972 100644 --- a/src/Helium/Instructions/Instances/Barrett.agda +++ b/src/Helium/Instructions/Instances/Barrett.agda @@ -22,11 +22,11 @@ open import Helium.Instructions.Core -- | z | < 2 ^ 31 -- Computes: -- z mod n -barret : (m -n : Expression [] (bits 32)) (t z : VecReg) (im : GenReg) → Procedure [] +barret : (m -n : Expression State [] (bits 32)) (t z : VecReg) (im : GenReg) → Procedure State [] barret m -n t z im = - index R (lit (im ′f)) ≔ m ∙ + *index R (lit im) ≔ m ∙ invoke vqrdmulh-s32,t,z,m [] ∙ - index R (lit (im ′f)) ≔ -n ∙ + *index R (lit im) ≔ -n ∙ invoke vmla-s32,z,t,-n [] ∙end where vqrdmulh-s32,t,z,m = diff --git a/src/Helium/Semantics/Axiomatic.agda b/src/Helium/Semantics/Axiomatic.agda index 2fa3db1..02e2a69 100644 --- a/src/Helium/Semantics/Axiomatic.agda +++ b/src/Helium/Semantics/Axiomatic.agda @@ -17,31 +17,38 @@ open import Helium.Data.Pseudocode.Algebra.Properties pseudocode open import Data.Nat using (ℕ) import Data.Unit -open import Data.Vec using (Vec) open import Function using (_∘_) -open import Helium.Data.Pseudocode.Core -import Helium.Semantics.Axiomatic.Core rawPseudocode as Core -import Helium.Semantics.Axiomatic.Assertion rawPseudocode as Assertion -import Helium.Semantics.Axiomatic.Term rawPseudocode as Term -import Helium.Semantics.Axiomatic.Triple rawPseudocode as Triple +import Helium.Semantics.Core rawPseudocode as Core′ +import Helium.Semantics.Axiomatic.Term rawPseudocode as Term′ +import Helium.Semantics.Axiomatic.Assertion rawPseudocode as Assertion′ -open Assertion.Construct public -open Assertion.Assertion public +private + proof-2≉0 : Core′.2≉0 + proof-2≉0 = ℝ.<⇒≉ (ℝ.n≢0∧x>0⇒n×x>0 2 (ℝ.≤∧≉⇒< ℝ.0≤1 (ℝ.1≉0 ∘ ℝ.Eq.sym))) ∘ ℝ.Eq.sym -open Assertion public - using (Assertion) +module Core where + open Core′ public hiding (shift) -open Term.Term public -open Term public - using (Term) + shift : ℤ → ℕ → ℤ + shift = Core′.shift proof-2≉0 + +open Core public using (⟦_⟧ₜ; ⟦_⟧ₜ′; Κ[_]_; 2≉0) -2≉0 : 2 ℝ.× 1ℝ ℝ.≉ 0ℝ -2≉0 = ℝ.<⇒≉ (ℝ.n≢0∧x>0⇒n×x>0 2 (ℝ.≤∧≉⇒< ℝ.0≤1 (ℝ.1≉0 ∘ ℝ.Eq.sym))) ∘ ℝ.Eq.sym +module Term where + open Term′ public hiding (module Semantics) + module Semantics {i} {j} {k} where + open Term′.Semantics {i} {j} {k} proof-2≉0 public -HoareTriple : ∀ {o} {Σ : Vec Type o} {n} {Γ : Vec Type n} {m} {Δ : Vec Type m} → Assertion Σ Γ Δ → Code.Statement Σ Γ → Assertion Σ Γ Δ → Set _ -HoareTriple = Triple.HoareTriple 2≉0 +open Term public using (Term; ↓_) hiding (module Term) +open Term.Term public + +module Assertion where + open Assertion′ public hiding (module Semantics) + module Semantics where + open Assertion′.Semantics proof-2≉0 public -ℰ : ∀ {o} {Σ : Vec Type o} {n} {Γ : Vec Type n} {m} {Δ : Vec Type m} {t : Type} → Code.Expression Σ Γ t → Term Σ Γ Δ t -ℰ = Term.ℰ 2≉0 +open Assertion public using (Assertion) hiding (module Assertion) +open Assertion.Assertion public +open Assertion.Construct public -open Triple.HoareTriple 2≉0 public +open import Helium.Semantics.Axiomatic.Triple rawPseudocode proof-2≉0 public diff --git a/src/Helium/Semantics/Axiomatic/Assertion.agda b/src/Helium/Semantics/Axiomatic/Assertion.agda index ab786e5..505abd8 100644 --- a/src/Helium/Semantics/Axiomatic/Assertion.agda +++ b/src/Helium/Semantics/Axiomatic/Assertion.agda @@ -15,115 +15,39 @@ module Helium.Semantics.Axiomatic.Assertion open RawPseudocode rawPseudocode -open import Data.Bool as Bool using (Bool) +import Data.Bool as Bool open import Data.Empty.Polymorphic using (⊥) -open import Data.Fin as Fin using (suc) +open import Data.Fin using (suc) open import Data.Fin.Patterns open import Data.Nat using (ℕ; suc) import Data.Nat.Properties as ℕₚ -open import Data.Product using (∃; _×_; _,_; proj₁; proj₂) +open import Data.Product using (∃; _×_; _,_; uncurry) open import Data.Sum using (_⊎_) open import Data.Unit.Polymorphic using (⊤) -open import Data.Vec as Vec using (Vec; []; _∷_; _++_) -open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -open import Function using (_$_) +open import Data.Vec as Vec using (Vec; []; _∷_; _++_; insert) +open import Data.Vec.Relation.Unary.All using (All; map) +import Data.Vec.Recursive as Vecᵣ +open import Function open import Helium.Data.Pseudocode.Core -open import Helium.Semantics.Axiomatic.Core rawPseudocode +open import Helium.Semantics.Core rawPseudocode open import Helium.Semantics.Axiomatic.Term rawPseudocode as Term using (Term) -open import Level using (_⊔_; Lift; lift; lower) renaming (suc to ℓsuc) +open import Level as L using (Lift; lift; lower) private variable t t′ : Type - m n o : ℕ + i j k m n o : ℕ Γ Δ Σ ts : Vec Type m -open Term.Term + ℓ = b₁ L.⊔ i₁ L.⊔ r₁ -data Assertion (Σ : Vec Type o) (Γ : Vec Type n) (Δ : Vec Type m) : Set (ℓsuc (b₁ ⊔ i₁ ⊔ r₁)) +open Term.Term -data Assertion Σ Γ Δ where +data Assertion (Σ : Vec Type o) (Γ : Vec Type n) (Δ : Vec Type m) : Set (L.suc ℓ) where all : Assertion Σ Γ (t ∷ Δ) → Assertion Σ Γ Δ some : Assertion Σ Γ (t ∷ Δ) → Assertion Σ Γ Δ pred : Term Σ Γ Δ bool → Assertion Σ Γ Δ - comb : ∀ {n} → (Vec (Set (b₁ ⊔ i₁ ⊔ r₁)) n → Set (b₁ ⊔ i₁ ⊔ r₁)) → Vec (Assertion Σ Γ Δ) n → Assertion Σ Γ Δ - -substVars : Assertion Σ Γ Δ → All (Term Σ ts Δ) Γ → Assertion Σ ts Δ -substVars (all P) ts = all (substVars P (Term.wknMeta′ ts)) -substVars (some P) ts = some (substVars P (Term.wknMeta′ ts)) -substVars (pred p) ts = pred (Term.substVars p ts) -substVars (comb f Ps) ts = comb f (helper Ps ts) - where - helper : ∀ {n m ts} → Vec (Assertion Σ _ Δ) n → All (Term {n = m} Σ ts Δ) Γ → Vec (Assertion Σ ts Δ) n - helper [] ts = [] - helper (P ∷ Ps) ts = substVars P ts ∷ helper Ps ts - -elimVar : Assertion Σ (t ∷ Γ) Δ → Term Σ Γ Δ t → Assertion Σ Γ Δ -elimVar (all P) t = all (elimVar P (Term.wknMeta t)) -elimVar (some P) t = some (elimVar P (Term.wknMeta t)) -elimVar (pred p) t = pred (Term.elimVar p t) -elimVar (comb f Ps) t = comb f (helper Ps t) - where - helper : ∀ {n} → Vec (Assertion Σ (_ ∷ Γ) Δ) n → Term Σ Γ Δ _ → Vec (Assertion Σ Γ Δ) n - helper [] t = [] - helper (P ∷ Ps) t = elimVar P t ∷ helper Ps t - -wknVar : Assertion Σ Γ Δ → Assertion Σ (t ∷ Γ) Δ -wknVar (all P) = all (wknVar P) -wknVar (some P) = some (wknVar P) -wknVar (pred p) = pred (Term.wknVar p) -wknVar (comb f Ps) = comb f (helper Ps) - where - helper : ∀ {n} → Vec (Assertion Σ Γ Δ) n → Vec (Assertion Σ (_ ∷ Γ) Δ) n - helper [] = [] - helper (P ∷ Ps) = wknVar P ∷ helper Ps - -wknVars : (ts : Vec Type n) → Assertion Σ Γ Δ → Assertion Σ (ts ++ Γ) Δ -wknVars τs (all P) = all (wknVars τs P) -wknVars τs (some P) = some (wknVars τs P) -wknVars τs (pred p) = pred (Term.wknVars τs p) -wknVars τs (comb f Ps) = comb f (helper Ps) - where - helper : ∀ {n} → Vec (Assertion Σ Γ Δ) n → Vec (Assertion Σ (τs ++ Γ) Δ) n - helper [] = [] - helper (P ∷ Ps) = wknVars τs P ∷ helper Ps - -addVars : Assertion Σ [] Δ → Assertion Σ Γ Δ -addVars (all P) = all (addVars P) -addVars (some P) = some (addVars P) -addVars (pred p) = pred (Term.addVars p) -addVars (comb f Ps) = comb f (helper Ps) - where - helper : ∀ {n} → Vec (Assertion Σ [] Δ) n → Vec (Assertion Σ Γ Δ) n - helper [] = [] - helper (P ∷ Ps) = addVars P ∷ helper Ps - -wknMetaAt : ∀ i → Assertion Σ Γ Δ → Assertion Σ Γ (Vec.insert Δ i t) -wknMetaAt i (all P) = all (wknMetaAt (suc i) P) -wknMetaAt i (some P) = some (wknMetaAt (suc i) P) -wknMetaAt i (pred p) = pred (Term.wknMetaAt i p) -wknMetaAt i (comb f Ps) = comb f (helper i Ps) - where - helper : ∀ {n} i → Vec (Assertion Σ Γ Δ) n → Vec (Assertion Σ Γ (Vec.insert Δ i _)) n - helper i [] = [] - helper i (P ∷ Ps) = wknMetaAt i P ∷ helper i Ps - --- NOTE: better to induct on P instead of ts? -wknMetas : (ts : Vec Type n) → Assertion Σ Γ Δ → Assertion Σ Γ (ts ++ Δ) -wknMetas [] P = P -wknMetas (_ ∷ ts) P = wknMetaAt 0F (wknMetas ts P) - -module _ (2≉0 : Term.2≉0) where - -- NOTE: better to induct on e here than in Term? - subst : Assertion Σ Γ Δ → {e : Code.Expression Σ Γ t} → Code.CanAssign Σ e → Term Σ Γ Δ t → Assertion Σ Γ Δ - subst (all P) e t = all (subst P e (Term.wknMeta t)) - subst (some P) e t = some (subst P e (Term.wknMeta t)) - subst (pred p) e t = pred (Term.subst 2≉0 p e t) - subst (comb f Ps) e t = comb f (helper Ps e t) - where - helper : ∀ {t n} → Vec (Assertion Σ Γ Δ) n → {e : Code.Expression Σ Γ t} → Code.CanAssign Σ e → Term Σ Γ Δ t → Vec (Assertion Σ Γ Δ) n - helper [] e t = [] - helper (P ∷ Ps) e t = subst P e t ∷ helper Ps e t + comb : (Set ℓ Vecᵣ.^ k → Set ℓ) → Vec (Assertion Σ Γ Δ) k → Assertion Σ Γ Δ module Construct where infixl 6 _∧_ @@ -136,38 +60,131 @@ module Construct where false = comb (λ _ → ⊥) [] _∧_ : Assertion Σ Γ Δ → Assertion Σ Γ Δ → Assertion Σ Γ Δ - P ∧ Q = comb (λ { (P ∷ Q ∷ []) → P × Q }) (P ∷ Q ∷ []) + P ∧ Q = comb (uncurry _×_) (P ∷ Q ∷ []) _∨_ : Assertion Σ Γ Δ → Assertion Σ Γ Δ → Assertion Σ Γ Δ - P ∨ Q = comb (λ { (P ∷ Q ∷ []) → P ⊎ Q }) (P ∷ Q ∷ []) + P ∨ Q = comb (uncurry _⊎_) (P ∷ Q ∷ []) equal : Term Σ Γ Δ t → Term Σ Γ Δ t → Assertion Σ Γ Δ - equal {t = bool} x y = pred Term.[ bool ][ x ≟ y ] - equal {t = int} x y = pred Term.[ int ][ x ≟ y ] - equal {t = fin n} x y = pred Term.[ fin ][ x ≟ y ] - equal {t = real} x y = pred Term.[ real ][ x ≟ y ] - equal {t = bit} x y = pred Term.[ bit ][ x ≟ y ] - equal {t = bits n} x y = pred Term.[ bits ][ x ≟ y ] - equal {t = tuple _ []} x y = true - equal {t = tuple _ (t ∷ ts)} x y = equal {t = t} (Term.func₁ proj₁ x) (Term.func₁ proj₁ y) ∧ equal (Term.func₁ proj₂ x) (Term.func₁ proj₂ y) - equal {t = array t 0} x y = true - equal {t = array t (suc n)} x y = all (equal {t = t} (index x) (index y)) + equal {t = bool} x y = pred (x ≟ y) + equal {t = int} x y = pred (x ≟ y) + equal {t = fin n} x y = pred (x ≟ y) + equal {t = real} x y = pred (x ≟ y) + equal {t = bit} x y = pred (x ≟ y) + equal {t = tuple []} x y = true + equal {t = tuple (t ∷ [])} x y = equal (head x) (head y) + equal {t = tuple (t ∷ t₁ ∷ ts)} x y = equal (head x) (head y) ∧ equal (tail x) (tail y) + equal {t = array t 0} x y = true + equal {t = array t (suc n)} x y = all (equal (index x) (index y)) where - index = λ v → Term.unbox (array t) $ - Term.func₁ proj₁ $ - Term.cut (array t) - (Term.cast (array t) (ℕₚ.+-comm 1 n) (Term.wknMeta v)) - (meta 0F) + index : Term Σ Γ Δ (array t (suc n)) → Term Σ Γ (fin (suc n) ∷ Δ) t + index t = unbox (slice (cast (ℕₚ.+-comm 1 _) (Term.Meta.weaken 0F t)) (meta 0F)) open Construct public -⟦_⟧ : Assertion Σ Γ Δ → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → Set (b₁ ⊔ i₁ ⊔ r₁) -⟦_⟧′ : Vec (Assertion Σ Γ Δ) n → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → Vec (Set (b₁ ⊔ i₁ ⊔ r₁)) n +module Var where + weaken : ∀ i → Assertion Σ Γ Δ → Assertion Σ (insert Γ i t) Δ + weaken i (all P) = all (weaken i P) + weaken i (some P) = some (weaken i P) + weaken i (pred p) = pred (Term.Var.weaken i p) + weaken i (comb f Ps) = comb f (helper i Ps) + where + helper : ∀ i → Vec (Assertion Σ Γ Δ) k → Vec (Assertion Σ (insert Γ i t) Δ) k + helper i [] = [] + helper i (P ∷ Ps) = weaken i P ∷ helper i Ps + + weakenAll : Assertion Σ [] Δ → Assertion Σ Γ Δ + weakenAll (all P) = all (weakenAll P) + weakenAll (some P) = some (weakenAll P) + weakenAll (pred p) = pred (Term.Var.weakenAll p) + weakenAll (comb f Ps) = comb f (helper Ps) + where + helper : Vec (Assertion Σ [] Δ) k → Vec (Assertion Σ Γ Δ) k + helper [] = [] + helper (P ∷ Ps) = weakenAll P ∷ helper Ps + + inject : ∀ (ts : Vec Type n) → Assertion Σ Γ Δ → Assertion Σ (Γ ++ ts) Δ + inject ts (all P) = all (inject ts P) + inject ts (some P) = some (inject ts P) + inject ts (pred p) = pred (Term.Var.inject ts p) + inject ts (comb f Ps) = comb f (helper ts Ps) + where + helper : ∀ (ts : Vec Type n) → Vec (Assertion Σ Γ Δ) k → Vec (Assertion Σ (Γ ++ ts) Δ) k + helper ts [] = [] + helper ts (P ∷ Ps) = inject ts P ∷ helper ts Ps + + raise : ∀ (ts : Vec Type n) → Assertion Σ Γ Δ → Assertion Σ (ts ++ Γ) Δ + raise ts (all P) = all (raise ts P) + raise ts (some P) = some (raise ts P) + raise ts (pred p) = pred (Term.Var.raise ts p) + raise ts (comb f Ps) = comb f (helper ts Ps) + where + helper : ∀ (ts : Vec Type n) → Vec (Assertion Σ Γ Δ) k → Vec (Assertion Σ (ts ++ Γ) Δ) k + helper ts [] = [] + helper ts (P ∷ Ps) = raise ts P ∷ helper ts Ps + + elim : ∀ i → Assertion Σ (insert Γ i t) Δ → Term Σ Γ Δ t → Assertion Σ Γ Δ + elim i (all P) e = all (elim i P (Term.Meta.weaken 0F e)) + elim i (some P) e = some (elim i P (Term.Meta.weaken 0F e)) + elim i (pred p) e = pred (Term.Var.elim i p e) + elim i (comb f Ps) e = comb f (helper i Ps e) + where + helper : ∀ i → Vec (Assertion Σ (insert Γ i t) Δ) k → Term Σ Γ Δ t → Vec (Assertion Σ Γ Δ) k + helper i [] e = [] + helper i (P ∷ Ps) e = elim i P e ∷ helper i Ps e + + elimAll : Assertion Σ Γ Δ → All (Term Σ ts Δ) Γ → Assertion Σ ts Δ + elimAll (all P) es = all (elimAll P (map (Term.Meta.weaken 0F) es)) + elimAll (some P) es = some (elimAll P (map (Term.Meta.weaken 0F) es)) + elimAll (pred p) es = pred (Term.Var.elimAll p es) + elimAll (comb f Ps) es = comb f (helper Ps es) + where + helper : Vec (Assertion Σ Γ Δ) n → All (Term Σ ts Δ) Γ → Vec (Assertion Σ ts Δ) n + helper [] es = [] + helper (P ∷ Ps) es = elimAll P es ∷ helper Ps es + +module Meta where + weaken : ∀ i → Assertion Σ Γ Δ → Assertion Σ Γ (insert Δ i t) + weaken i (all P) = all (weaken (suc i) P) + weaken i (some P) = some (weaken (suc i) P) + weaken i (pred p) = pred (Term.Meta.weaken i p) + weaken i (comb f Ps) = comb f (helper i Ps) + where + helper : ∀ i → Vec (Assertion Σ Γ Δ) k → Vec (Assertion Σ Γ (insert Δ i t)) k + helper i [] = [] + helper i (P ∷ Ps) = weaken i P ∷ helper i Ps + + elim : ∀ i → Assertion Σ Γ (insert Δ i t) → Term Σ Γ Δ t → Assertion Σ Γ Δ + elim i (all P) e = all (elim (suc i) P (Term.Meta.weaken 0F e)) + elim i (some P) e = some (elim (suc i) P (Term.Meta.weaken 0F e)) + elim i (pred p) e = pred (Term.Meta.elim i p e) + elim i (comb f Ps) e = comb f (helper i Ps e) + where + helper : ∀ i → Vec (Assertion Σ Γ (insert Δ i t)) k → Term Σ Γ Δ t → Vec (Assertion Σ Γ Δ) k + helper i [] e = [] + helper i (P ∷ Ps) e = elim i P e ∷ helper i Ps e + +subst : Assertion Σ Γ Δ → Reference Σ Γ t → Term Σ Γ Δ t → Assertion Σ Γ Δ +subst (all P) ref val = all (subst P ref (Term.Meta.weaken 0F val)) +subst (some P) ref val = some (subst P ref (Term.Meta.weaken 0F val)) +subst (pred p) ref val = pred (Term.subst p ref val) +subst (comb f Ps) ref val = comb f (helper Ps ref val) + where + helper : Vec (Assertion Σ Γ Δ) k → Reference Σ Γ t → Term Σ Γ Δ t → Vec (Assertion Σ Γ Δ) k + helper [] ref val = [] + helper (P ∷ Ps) ref val = subst P ref val ∷ Ps + + +module Semantics (2≉0 : 2≉0) where + module TS {i} {j} {k} = Term.Semantics {i} {j} {k} 2≉0 + + ⟦_⟧ : Assertion Σ Γ Δ → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → Set ℓ + ⟦_⟧′ : Vec (Assertion Σ Γ Δ) n → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → Vec (Set ℓ) n -⟦ all P ⟧ σ γ δ = ∀ x → ⟦ P ⟧ σ γ (x , δ) -⟦ some P ⟧ σ γ δ = ∃ λ x → ⟦ P ⟧ σ γ (x , δ) -⟦ pred p ⟧ σ γ δ = Lift (b₁ ⊔ i₁ ⊔ r₁) (Bool.T (lower (Term.⟦ p ⟧ σ γ δ))) -⟦ comb f Ps ⟧ σ γ δ = f (⟦ Ps ⟧′ σ γ δ) + ⟦_⟧ {Δ = Δ} (all P) σ γ δ = ∀ x → ⟦ P ⟧ σ γ (cons′ Δ x δ) + ⟦_⟧ {Δ = Δ} (some P) σ γ δ = ∃ λ x → ⟦ P ⟧ σ γ (cons′ Δ x δ) + ⟦ pred p ⟧ σ γ δ = Lift ℓ (Bool.T (lower (TS.⟦ p ⟧ σ γ δ))) + ⟦ comb f Ps ⟧ σ γ δ = f (Vecᵣ.fromVec (⟦ Ps ⟧′ σ γ δ)) -⟦ [] ⟧′ σ γ δ = [] -⟦ P ∷ Ps ⟧′ σ γ δ = ⟦ P ⟧ σ γ δ ∷ ⟦ Ps ⟧′ σ γ δ + ⟦ [] ⟧′ σ γ δ = [] + ⟦ P ∷ Ps ⟧′ σ γ δ = ⟦ P ⟧ σ γ δ ∷ ⟦ Ps ⟧′ σ γ δ diff --git a/src/Helium/Semantics/Axiomatic/Core.agda b/src/Helium/Semantics/Axiomatic/Core.agda deleted file mode 100644 index a65a6d0..0000000 --- a/src/Helium/Semantics/Axiomatic/Core.agda +++ /dev/null @@ -1,85 +0,0 @@ ------------------------------------------------------------------------- --- Agda Helium --- --- Base definitions for the axiomatic semantics ------------------------------------------------------------------------- - -{-# OPTIONS --safe --without-K #-} - -open import Helium.Data.Pseudocode.Algebra using (RawPseudocode) - -module Helium.Semantics.Axiomatic.Core - {b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃} - (rawPseudocode : RawPseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) - where - -private - open module C = RawPseudocode rawPseudocode - -open import Data.Bool as Bool using (Bool) -open import Data.Fin as Fin using (Fin; zero; suc) -open import Data.Fin.Patterns -open import Data.Nat as ℕ using (ℕ; suc) -import Data.Nat.Induction as Natᵢ -import Data.Nat.Properties as ℕₚ -open import Data.Product as × using (_×_; _,_; uncurry) -open import Data.Sum using (_⊎_) -open import Data.Unit using (⊤) -open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup) -open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -open import Function using (_on_) -open import Helium.Data.Pseudocode.Core -open import Helium.Data.Pseudocode.Properties -import Induction.WellFounded as Wf -open import Level using (_⊔_; Lift; lift) -import Relation.Binary.Construct.On as On -open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; cong₂) -open import Relation.Nullary using (Dec; does; yes; no) -open import Relation.Nullary.Decidable.Core using (True; toWitness; map′) -open import Relation.Nullary.Product using (_×-dec_) -open import Relation.Unary using (_⊆_) - -private - variable - t t′ : Type - m n : ℕ - Γ Δ Σ ts : Vec Type m - -⟦_⟧ₜ : Type → Set (b₁ ⊔ i₁ ⊔ r₁) -⟦_⟧ₜ′ : Vec Type n → Set (b₁ ⊔ i₁ ⊔ r₁) - -⟦ bool ⟧ₜ = Lift (b₁ ⊔ i₁ ⊔ r₁) Bool -⟦ int ⟧ₜ = Lift (b₁ ⊔ r₁) ℤ -⟦ fin n ⟧ₜ = Lift (b₁ ⊔ i₁ ⊔ r₁) (Fin n) -⟦ real ⟧ₜ = Lift (b₁ ⊔ i₁) ℝ -⟦ bit ⟧ₜ = Lift (i₁ ⊔ r₁) Bit -⟦ bits n ⟧ₜ = Vec ⟦ bit ⟧ₜ n -⟦ tuple n ts ⟧ₜ = ⟦ ts ⟧ₜ′ -⟦ array t n ⟧ₜ = Vec ⟦ t ⟧ₜ n - -⟦ [] ⟧ₜ′ = Lift (b₁ ⊔ i₁ ⊔ r₁) ⊤ -⟦ t ∷ ts ⟧ₜ′ = ⟦ t ⟧ₜ × ⟦ ts ⟧ₜ′ - -fetch : ∀ i → ⟦ Γ ⟧ₜ′ → ⟦ lookup Γ i ⟧ₜ -fetch {Γ = _ ∷ _} 0F (x , _) = x -fetch {Γ = _ ∷ _} (suc i) (_ , xs) = fetch i xs - -Transform : Vec Type m → Type → Set (b₁ ⊔ i₁ ⊔ r₁) -Transform ts t = ⟦ ts ⟧ₜ′ → ⟦ t ⟧ₜ - -Predicate : Vec Type m → Set (b₁ ⊔ i₁ ⊔ r₁) -Predicate ts = ⟦ ts ⟧ₜ′ → Bool - --- data HoareTriple {n Γ m Δ} : Assertion Σ {n} Γ {m} Δ → Statement Γ → Assertion Σ Γ Δ → Set (b₁ ⊔ i₁ ⊔ r₁) where --- _∙_ : ∀ {P Q R s₁ s₂} → HoareTriple P s₁ Q → HoareTriple Q s₂ R → HoareTriple P (s₁ ∙ s₂) R --- skip : ∀ {P} → HoareTriple P skip P - --- assign : ∀ {P t ref canAssign e} → HoareTriple (asrtSubst P (toWitness canAssign) (ℰ e)) (_≔_ {t = t} ref {canAssign} e) P --- declare : ∀ {t P Q e s} → HoareTriple (P ∧ equal (var 0F) (termWknVar (ℰ e))) s (asrtWknVar Q) → HoareTriple (asrtElimVar P (ℰ e)) (declare {t = t} e s) Q --- invoke : ∀ {m ts P Q s e} → HoareTriple (assignMetas Δ ts (All.tabulate var) (asrtAddVars P)) s (asrtAddVars Q) → HoareTriple (assignMetas Δ ts (All.tabulate (λ i → ℰ (All.lookup i e))) (asrtAddVars P)) (invoke {m = m} {ts} (s ∙end) e) (asrtAddVars Q) --- if : ∀ {P Q e s₁ s₂} → HoareTriple (P ∧ equal (ℰ e) (𝒦 (Bool.true ′b))) s₁ Q → HoareTriple (P ∧ equal (ℰ e) (𝒦 (Bool.false ′b))) s₂ Q → HoareTriple P (if e then s₁ else s₂) Q --- for : ∀ {m} {I : Assertion Σ Γ (fin (suc m) ∷ Δ)} {s} → HoareTriple (some (asrtWknVar (asrtWknMetaAt 1F I) ∧ equal (meta 1F) (var 0F) ∧ equal (meta 0F) (func (λ { _ (lift x ∷ []) → lift (Fin.inject₁ x) }) (meta 1F ∷ [])))) s (some (asrtWknVar (asrtWknMetaAt 1F I) ∧ equal (meta 0F) (func (λ { _ (lift x ∷ []) → lift (Fin.suc x) }) (meta 1F ∷ [])))) → HoareTriple (some (I ∧ equal (meta 0F) (func (λ _ _ → lift 0F) []))) (for m s) (some (I ∧ equal (meta 0F) (func (λ _ _ → lift (Fin.fromℕ m)) []))) - --- consequence : ∀ {P₁ P₂ Q₁ Q₂ s} → ⟦ P₁ ⟧ₐ ⊆ ⟦ P₂ ⟧ₐ → HoareTriple P₂ s Q₂ → ⟦ Q₂ ⟧ₐ ⊆ ⟦ Q₁ ⟧ₐ → HoareTriple P₁ s Q₁ --- some : ∀ {t P Q s} → HoareTriple P s Q → HoareTriple (some {t = t} P) s (some Q) --- frame : ∀ {P Q R s} → HoareTriple P s Q → HoareTriple (P * R) s (Q * R) diff --git a/src/Helium/Semantics/Axiomatic/Term.agda b/src/Helium/Semantics/Axiomatic/Term.agda index c9ddd02..08eac5f 100644 --- a/src/Helium/Semantics/Axiomatic/Term.agda +++ b/src/Helium/Semantics/Axiomatic/Term.agda @@ -13,373 +13,530 @@ module Helium.Semantics.Axiomatic.Term (rawPseudocode : RawPseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) where + open RawPseudocode rawPseudocode import Data.Bool as Bool -open import Data.Fin as Fin using (Fin; suc) -import Data.Fin.Properties as Finₚ +open import Data.Empty using (⊥-elim) +open import Data.Fin as Fin using (Fin; suc; punchOut) open import Data.Fin.Patterns -open import Data.Nat as ℕ using (ℕ; suc) +import Data.Integer as 𝕀 +import Data.Fin.Properties as Finₚ +open import Data.Nat as ℕ using (ℕ; suc; _≤_; z≤n; s≤s; _⊔_) import Data.Nat.Properties as ℕₚ -open import Data.Product using (∃; _×_; _,_; proj₁; proj₂; uncurry; dmap) -open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup) -import Data.Vec.Functional as VecF +open import Data.Product using (∃; _,_; dmap) +open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup; insert; remove; map; zipWith; take; drop) import Data.Vec.Properties as Vecₚ +open import Data.Vec.Recursive as Vecᵣ using (2+_) open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) -open import Function using (_∘_; _∘₂_; id; flip) +open import Function open import Helium.Data.Pseudocode.Core -import Helium.Data.Pseudocode.Manipulate as M -open import Helium.Semantics.Axiomatic.Core rawPseudocode -open import Level using (_⊔_; lift; lower) -open import Relation.Binary.PropositionalEquality hiding ([_]) renaming (subst to ≡-subst) -open import Relation.Nullary using (¬_; does; yes; no) -open import Relation.Nullary.Decidable.Core using (True; toWitness) -open import Relation.Nullary.Negation using (contradiction) +open import Helium.Data.Pseudocode.Manipulate hiding (module Cast) +open import Helium.Semantics.Core rawPseudocode +open import Level as L using (lift; lower) +open import Relation.Binary.PropositionalEquality hiding (subst) +open import Relation.Nullary using (does; yes; no) private variable - t t′ t₁ t₂ : Type - m n o : ℕ - Γ Δ Σ ts : Vec Type m - -data Term (Σ : Vec Type o) (Γ : Vec Type n) (Δ : Vec Type m) : Type → Set (b₁ ⊔ i₁ ⊔ r₁) where - state : ∀ i → Term Σ Γ Δ (lookup Σ i) - var : ∀ i → Term Σ Γ Δ (lookup Γ i) - meta : ∀ i → Term Σ Γ Δ (lookup Δ i) - func : Transform ts t → All (Term Σ Γ Δ) ts → Term Σ Γ Δ t - -castType : Term Σ Γ Δ t → t ≡ t′ → Term Σ Γ Δ t′ -castType (state i) refl = state i -castType (var i) refl = var i -castType (meta i) refl = meta i -castType (func f ts) eq = func (≡-subst (Transform _) eq f) ts - -substState : Term Σ Γ Δ t → ∀ i → Term Σ Γ Δ (lookup Σ i) → Term Σ Γ Δ t -substState (state i) j t′ with i Fin.≟ j -... | yes refl = t′ -... | no _ = state i -substState (var i) j t′ = var i -substState (meta i) j t′ = meta i -substState (func f ts) j t′ = func f (helper ts j t′) - where - helper : ∀ {n ts} → All (Term Σ Γ Δ) {n} ts → ∀ i → Term Σ Γ Δ (lookup Σ i) → All (Term Σ Γ Δ) ts - helper [] i t′ = [] - helper (t ∷ ts) i t′ = substState t i t′ ∷ helper ts i t′ - -substVar : Term Σ Γ Δ t → ∀ i → Term Σ Γ Δ (lookup Γ i) → Term Σ Γ Δ t -substVar (state i) j t′ = state i -substVar (var i) j t′ with i Fin.≟ j -... | yes refl = t′ -... | no _ = var i -substVar (meta i) j t′ = meta i -substVar (func f ts) j t′ = func f (helper ts j t′) - where - helper : ∀ {n ts} → All (Term Σ Γ Δ) {n} ts → ∀ i → Term Σ Γ Δ (lookup Γ i) → All (Term Σ Γ Δ) ts - helper [] i t′ = [] - helper (t ∷ ts) i t′ = substVar t i t′ ∷ helper ts i t′ - -substVars : Term Σ Γ Δ t → All (Term Σ ts Δ) Γ → Term Σ ts Δ t -substVars (state i) ts = state i -substVars (var i) ts = All.lookup i ts -substVars (meta i) ts = meta i -substVars (func f ts′) ts = func f (helper ts′ ts) - where - helper : ∀ {ts ts′} → All (Term Σ Γ Δ) {n} ts → All (Term {n = m} Σ ts′ Δ) Γ → All (Term Σ ts′ Δ) ts - helper [] ts = [] - helper (t ∷ ts′) ts = substVars t ts ∷ helper ts′ ts - -elimVar : Term Σ (t′ ∷ Γ) Δ t → Term Σ Γ Δ t′ → Term Σ Γ Δ t -elimVar (state i) t′ = state i -elimVar (var 0F) t′ = t′ -elimVar (var (suc i)) t′ = var i -elimVar (meta i) t′ = meta i -elimVar (func f ts) t′ = func f (helper ts t′) - where - helper : ∀ {n ts} → All (Term Σ (_ ∷ Γ) Δ) {n} ts → Term Σ Γ Δ _ → All (Term Σ Γ Δ) ts - helper [] t′ = [] - helper (t ∷ ts) t′ = elimVar t t′ ∷ helper ts t′ - -wknVar : Term Σ Γ Δ t → Term Σ (t′ ∷ Γ) Δ t -wknVar (state i) = state i -wknVar (var i) = var (suc i) -wknVar (meta i) = meta i -wknVar (func f ts) = func f (helper ts) - where - helper : ∀ {n ts} → All (Term Σ Γ Δ) {n} ts → All (Term Σ (_ ∷ Γ) Δ) ts - helper [] = [] - helper (t ∷ ts) = wknVar t ∷ helper ts - -wknVars : (ts : Vec Type n) → Term Σ Γ Δ t → Term Σ (ts ++ Γ) Δ t -wknVars τs (state i) = state i -wknVars τs (var i) = castType (var (Fin.raise (Vec.length τs) i)) (Vecₚ.lookup-++ʳ τs _ i) -wknVars τs (meta i) = meta i -wknVars τs (func f ts) = func f (helper ts) - where - helper : ∀ {n ts} → All (Term Σ Γ Δ) {n} ts → All (Term Σ (τs ++ Γ) Δ) ts - helper [] = [] - helper (t ∷ ts) = wknVars τs t ∷ helper ts - -addVars : Term Σ [] Δ t → Term Σ Γ Δ t -addVars (state i) = state i -addVars (meta i) = meta i -addVars (func f ts) = func f (helper ts) - where - helper : ∀ {n ts} → All (Term Σ [] Δ) {n} ts → All (Term Σ Γ Δ) ts - helper [] = [] - helper (t ∷ ts) = addVars t ∷ helper ts - -wknMetaAt : ∀ i → Term Σ Γ Δ t → Term Σ Γ (Vec.insert Δ i t′) t -wknMetaAt′ : ∀ i → All (Term Σ Γ Δ) ts → All (Term Σ Γ (Vec.insert Δ i t′)) ts - -wknMetaAt i (state j) = state j -wknMetaAt i (var j) = var j -wknMetaAt i (meta j) = castType (meta (Fin.punchIn i j)) (Vecₚ.insert-punchIn _ i _ j) -wknMetaAt i (func f ts) = func f (wknMetaAt′ i ts) - -wknMetaAt′ i [] = [] -wknMetaAt′ i (t ∷ ts) = wknMetaAt i t ∷ wknMetaAt′ i ts - -wknMeta : Term Σ Γ Δ t → Term Σ Γ (t′ ∷ Δ) t -wknMeta = wknMetaAt 0F - -wknMeta′ : All (Term Σ Γ Δ) ts → All (Term Σ Γ (t′ ∷ Δ)) ts -wknMeta′ = wknMetaAt′ 0F - -wknMetas : (ts : Vec Type n) → Term Σ Γ Δ t → Term Σ Γ (ts ++ Δ) t -wknMetas τs (state i) = state i -wknMetas τs (var i) = var i -wknMetas τs (meta i) = castType (meta (Fin.raise (Vec.length τs) i)) (Vecₚ.lookup-++ʳ τs _ i) -wknMetas τs (func f ts) = func f (helper ts) - where - helper : ∀ {n ts} → All (Term Σ Γ Δ) {n} ts → All (Term Σ Γ (τs ++ Δ)) ts - helper [] = [] - helper (t ∷ ts) = wknMetas τs t ∷ helper ts - -func₀ : ⟦ t ⟧ₜ → Term Σ Γ Δ t -func₀ f = func (λ _ → f) [] - -func₁ : (⟦ t ⟧ₜ → ⟦ t′ ⟧ₜ) → Term Σ Γ Δ t → Term Σ Γ Δ t′ -func₁ f t = func (λ (x , _) → f x) (t ∷ []) - -func₂ : (⟦ t₁ ⟧ₜ → ⟦ t₂ ⟧ₜ → ⟦ t′ ⟧ₜ) → Term Σ Γ Δ t₁ → Term Σ Γ Δ t₂ → Term Σ Γ Δ t′ -func₂ f t₁ t₂ = func (λ (x , y , _) → f x y) (t₁ ∷ t₂ ∷ []) - -[_][_≟_] : HasEquality t → Term Σ Γ Δ t → Term Σ Γ Δ t → Term Σ Γ Δ bool -[ bool ][ t ≟ t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x Bool.≟ y))) t t′ -[ int ][ t ≟ t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x ≟ᶻ y))) t t′ -[ fin ][ t ≟ t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x Fin.≟ y))) t t′ -[ real ][ t ≟ t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x ≟ʳ y))) t t′ -[ bit ][ t ≟ t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x ≟ᵇ₁ y))) t t′ -[ bits ][ t ≟ t′ ] = func₂ (λ xs ys → lift (does (VecF.fromVec (Vec.map lower xs) ≟ᵇ VecF.fromVec (Vec.map lower ys)))) t t′ - -[_][_<?_] : IsNumeric t → Term Σ Γ Δ t → Term Σ Γ Δ t → Term Σ Γ Δ bool -[ int ][ t <? t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x <ᶻ? y))) t t′ -[ real ][ t <? t′ ] = func₂ (λ (lift x) (lift y) → lift (does (x <ʳ? y))) t t′ - -[_][_] : ∀ t → Term Σ Γ Δ (elemType t) → Term Σ Γ Δ (asType t 1) -[ bits ][ t ] = func₁ (_∷ []) t -[ array τ ][ t ] = func₁ (_∷ []) t - -unbox : ∀ t → Term Σ Γ Δ (asType t 1) → Term Σ Γ Δ (elemType t) -unbox bits = func₁ Vec.head -unbox (array t) = func₁ Vec.head - -castV : ∀ {a} {A : Set a} {m n} → .(eq : m ≡ n) → Vec A m → Vec A n -castV {n = 0} eq [] = [] -castV {n = suc n} eq (x ∷ xs) = x ∷ castV (ℕₚ.suc-injective eq) xs - -cast′ : ∀ t → .(eq : m ≡ n) → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ -cast′ bits eq = castV eq -cast′ (array τ) eq = castV eq - -cast : ∀ t → .(eq : m ≡ n) → Term Σ Γ Δ (asType t m) → Term Σ Γ Δ (asType t n) -cast τ eq = func₁ (cast′ τ eq) - -[_][-_] : IsNumeric t → Term Σ Γ Δ t → Term Σ Γ Δ t -[ int ][- t ] = func₁ (lift ∘ ℤ.-_ ∘ lower) t -[ real ][- t ] = func₁ (lift ∘ ℝ.-_ ∘ lower) t - -[_,_,_,_][_+_] : ∀ t₁ t₂ → (isNum₁ : True (isNumeric? t₁)) → (isNum₂ : True (isNumeric? t₂)) → Term Σ Γ Δ t₁ → Term Σ Γ Δ t₂ → Term Σ Γ Δ (combineNumeric t₁ t₂ isNum₁ isNum₂) -[ int , int , _ , _ ][ t + t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℤ.+ y)) t t′ -[ int , real , _ , _ ][ t + t′ ] = func₂ (λ (lift x) (lift y) → lift (x /1 ℝ.+ y)) t t′ -[ real , int , _ , _ ][ t + t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℝ.+ y /1)) t t′ -[ real , real , _ , _ ][ t + t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℝ.+ y)) t t′ - -[_,_,_,_][_*_] : ∀ t₁ t₂ → (isNum₁ : True (isNumeric? t₁)) → (isNum₂ : True (isNumeric? t₂)) → Term Σ Γ Δ t₁ → Term Σ Γ Δ t₂ → Term Σ Γ Δ (combineNumeric t₁ t₂ isNum₁ isNum₂) -[ int , int , _ , _ ][ t * t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℤ.* y)) t t′ -[ int , real , _ , _ ][ t * t′ ] = func₂ (λ (lift x) (lift y) → lift (x /1 ℝ.* y)) t t′ -[ real , int , _ , _ ][ t * t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℝ.* y /1)) t t′ -[ real , real , _ , _ ][ t * t′ ] = func₂ (λ (lift x) (lift y) → lift (x ℝ.* y)) t t′ - -[_][_^_] : IsNumeric t → Term Σ Γ Δ t → ℕ → Term Σ Γ Δ t -[ int ][ t ^ n ] = func₁ (lift ∘ (ℤ′._^′ n) ∘ lower) t -[ real ][ t ^ n ] = func₁ (lift ∘ (ℝ′._^′ n) ∘ lower) t - -2≉0 : Set _ -2≉0 = ¬ 2 ℝ′.×′ 1ℝ ℝ.≈ 0ℝ - -[_][_>>_] : 2≉0 → Term Σ Γ Δ int → ℕ → Term Σ Γ Δ int -[ 2≉0 ][ t >> n ] = func₁ (lift ∘ ⌊_⌋ ∘ (ℝ._* 2≉0 ℝ.⁻¹ ℝ′.^′ n) ∘ _/1 ∘ lower) t - --- 0 of y is 0 of result -join′ : ∀ t → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ → ⟦ asType t (n ℕ.+ m) ⟧ₜ -join′ bits = flip _++_ -join′ (array t) = flip _++_ - -take′ : ∀ t m {n} → ⟦ asType t (m ℕ.+ n) ⟧ₜ → ⟦ asType t m ⟧ₜ -take′ bits m = Vec.take m -take′ (array t) m = Vec.take m - -drop′ : ∀ t m {n} → ⟦ asType t (m ℕ.+ n) ⟧ₜ → ⟦ asType t n ⟧ₜ -drop′ bits m = Vec.drop m -drop′ (array t) m = Vec.drop m - -private - m≤n⇒m+k≡n : ∀ {m n} → m ℕ.≤ n → ∃ λ k → m ℕ.+ k ≡ n - m≤n⇒m+k≡n ℕ.z≤n = _ , refl - m≤n⇒m+k≡n (ℕ.s≤s m≤n) = dmap id (cong suc) (m≤n⇒m+k≡n m≤n) - - slicedSize : ∀ n m (i : Fin (suc n)) → ∃ λ k → n ℕ.+ m ≡ Fin.toℕ i ℕ.+ (m ℕ.+ k) × Fin.toℕ i ℕ.+ k ≡ n - slicedSize n m i = k , (begin - n ℕ.+ m ≡˘⟨ cong (ℕ._+ m) (proj₂ i+k≡n) ⟩ - (Fin.toℕ i ℕ.+ k) ℕ.+ m ≡⟨ ℕₚ.+-assoc (Fin.toℕ i) k m ⟩ - Fin.toℕ i ℕ.+ (k ℕ.+ m) ≡⟨ cong (Fin.toℕ i ℕ.+_) (ℕₚ.+-comm k m) ⟩ - Fin.toℕ i ℕ.+ (m ℕ.+ k) ∎) , - proj₂ i+k≡n - where - open ≡-Reasoning - i+k≡n = m≤n⇒m+k≡n (ℕₚ.≤-pred (Finₚ.toℕ<n i)) - k = proj₁ i+k≡n - --- 0 of x is i of result -splice′ : ∀ t → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ → ⟦ fin (suc n) ⟧ₜ → ⟦ asType t (n ℕ.+ m) ⟧ₜ -splice′ {m = m} t x y (lift i) = cast′ t eq (join′ t (join′ t high x) low) + t t′ t₁ t₂ : Type + i j k m n o : ℕ + Γ Δ Σ ts : Vec Type m + + ℓ = b₁ L.⊔ i₁ L.⊔ r₁ + + punchOut-insert : ∀ {a} {A : Set a} (xs : Vec A n) {i j} (i≢j : i ≢ j) x → lookup xs (punchOut i≢j) ≡ lookup (insert xs i x) j + punchOut-insert xs {i} {j} i≢j x = begin + lookup xs (punchOut i≢j) ≡˘⟨ cong (flip lookup (punchOut i≢j)) (Vecₚ.remove-insert xs i x) ⟩ + lookup (remove (insert xs i x) i) (punchOut i≢j) ≡⟨ Vecₚ.remove-punchOut (insert xs i x) i≢j ⟩ + lookup (insert xs i x) j ∎ + where open ≡-Reasoning + + open ℕₚ.≤-Reasoning + + ⨆[_]_ : ∀ n → ℕ Vecᵣ.^ n → ℕ + ⨆[_]_ = Vecᵣ.foldl (const ℕ) 0 id (const (flip ℕ._⊔_)) + + ⨆-step : ∀ m x xs → ⨆[ 2+ m ] (x , xs) ≡ x ⊔ ⨆[ suc m ] xs + ⨆-step 0 x xs = refl + ⨆-step (suc m) x (y , xs) = begin-equality + ⨆[ 2+ suc m ] (x , y , xs) ≡⟨ ⨆-step m (x ⊔ y) xs ⟩ + x ⊔ y ⊔ ⨆[ suc m ] xs ≡⟨ ℕₚ.⊔-assoc x y _ ⟩ + x ⊔ (y ⊔ ⨆[ suc m ] xs) ≡˘⟨ cong (_ ⊔_) (⨆-step m y xs) ⟩ + x ⊔ ⨆[ 2+ m ] (y , xs) ∎ + + lookup-⨆-≤ : ∀ i (xs : ℕ Vecᵣ.^ n) → Vecᵣ.lookup i xs ≤ ⨆[ n ] xs + lookup-⨆-≤ {1} 0F x = ℕₚ.≤-refl + lookup-⨆-≤ {2+ n} 0F (x , xs) = begin + x ≤⟨ ℕₚ.m≤m⊔n x _ ⟩ + x ⊔ ⨆[ suc n ] xs ≡˘⟨ ⨆-step n x xs ⟩ + ⨆[ 2+ n ] (x , xs) ∎ + lookup-⨆-≤ {2+ n} (suc i) (x , xs) = begin + Vecᵣ.lookup i xs ≤⟨ lookup-⨆-≤ i xs ⟩ + ⨆[ suc n ] xs ≤⟨ ℕₚ.m≤n⊔m x _ ⟩ + x ⊔ ⨆[ suc n ] xs ≡˘⟨ ⨆-step n x xs ⟩ + ⨆[ 2+ n ] (x , xs) ∎ + +data Term (Σ : Vec Type i) (Γ : Vec Type j) (Δ : Vec Type k) : Type → Set ℓ where + lit : ⟦ t ⟧ₜ → Term Σ Γ Δ t + state : ∀ i → Term Σ Γ Δ (lookup Σ i) + var : ∀ i → Term Σ Γ Δ (lookup Γ i) + meta : ∀ i → Term Σ Γ Δ (lookup Δ i) + _≟_ : ⦃ HasEquality t ⦄ → Term Σ Γ Δ t → Term Σ Γ Δ t → Term Σ Γ Δ bool + _<?_ : ⦃ Ordered t ⦄ → Term Σ Γ Δ t → Term Σ Γ Δ t → Term Σ Γ Δ bool + inv : Term Σ Γ Δ bool → Term Σ Γ Δ bool + _&&_ : Term Σ Γ Δ bool → Term Σ Γ Δ bool → Term Σ Γ Δ bool + _||_ : Term Σ Γ Δ bool → Term Σ Γ Δ bool → Term Σ Γ Δ bool + not : Term Σ Γ Δ (bits n) → Term Σ Γ Δ (bits n) + _and_ : Term Σ Γ Δ (bits n) → Term Σ Γ Δ (bits n) → Term Σ Γ Δ (bits n) + _or_ : Term Σ Γ Δ (bits n) → Term Σ Γ Δ (bits n) → Term Σ Γ Δ (bits n) + [_] : Term Σ Γ Δ t → Term Σ Γ Δ (array t 1) + unbox : Term Σ Γ Δ (array t 1) → Term Σ Γ Δ t + merge : Term Σ Γ Δ (array t m) → Term Σ Γ Δ (array t n) → Term Σ Γ Δ (fin (suc n)) → Term Σ Γ Δ (array t (n ℕ.+ m)) + slice : Term Σ Γ Δ (array t (n ℕ.+ m)) → Term Σ Γ Δ (fin (suc n)) → Term Σ Γ Δ (array t m) + cut : Term Σ Γ Δ (array t (n ℕ.+ m)) → Term Σ Γ Δ (fin (suc n)) → Term Σ Γ Δ (array t n) + cast : .(eq : m ≡ n) → Term Σ Γ Δ (array t m) → Term Σ Γ Δ (array t n) + -_ : ⦃ IsNumeric t ⦄ → Term Σ Γ Δ t → Term Σ Γ Δ t + _+_ : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → Term Σ Γ Δ t₁ → Term Σ Γ Δ t₂ → Term Σ Γ Δ (isNum₁ +ᵗ isNum₂) + _*_ : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → Term Σ Γ Δ t₁ → Term Σ Γ Δ t₂ → Term Σ Γ Δ (isNum₁ +ᵗ isNum₂) + _^_ : ⦃ IsNumeric t ⦄ → Term Σ Γ Δ t → ℕ → Term Σ Γ Δ t + _>>_ : Term Σ Γ Δ int → (n : ℕ) → Term Σ Γ Δ int + rnd : Term Σ Γ Δ real → Term Σ Γ Δ int + fin : ∀ {ms} (f : literalTypes (map fin ms) → Fin n) → Term Σ Γ Δ (tuple {n = o} (map fin ms)) → Term Σ Γ Δ (fin n) + asInt : Term Σ Γ Δ (fin n) → Term Σ Γ Δ int + nil : Term Σ Γ Δ (tuple []) + cons : Term Σ Γ Δ t → Term Σ Γ Δ (tuple ts) → Term Σ Γ Δ (tuple (t ∷ ts)) + head : Term Σ Γ Δ (tuple (t ∷ ts)) → Term Σ Γ Δ t + tail : Term Σ Γ Δ (tuple (t ∷ ts)) → Term Σ Γ Δ (tuple ts) + if_then_else_ : Term Σ Γ Δ bool → Term Σ Γ Δ t → Term Σ Γ Δ t → Term Σ Γ Δ t + +↓_ : Expression Σ Γ t → Term Σ Γ Δ t +↓ e = go (Flatten.expr e) (Flatten.expr-depth e) where - reasoning = slicedSize _ m i - k = proj₁ reasoning - n≡i+k = sym (proj₂ (proj₂ reasoning)) - low = take′ t (Fin.toℕ i) (cast′ t n≡i+k y) - high = drop′ t (Fin.toℕ i) (cast′ t n≡i+k y) - eq = sym (proj₁ (proj₂ reasoning)) - -splice : ∀ t → Term Σ Γ Δ (asType t m) → Term Σ Γ Δ (asType t n) → Term Σ Γ Δ (fin (suc n)) → Term Σ Γ Δ (asType t (n ℕ.+ m)) -splice τ t₁ t₂ t′ = func (λ (x , y , i , _) → splice′ τ x y i) (t₁ ∷ t₂ ∷ t′ ∷ []) - --- i of x is 0 of first -cut′ : ∀ t → ⟦ asType t (n ℕ.+ m) ⟧ₜ → ⟦ fin (suc n) ⟧ₜ → ⟦ asType t m ∷ asType t n ∷ [] ⟧ₜ′ -cut′ {m = m} t x (lift i) = middle , cast′ t i+k≡n (join′ t high low) , _ - where - reasoning = slicedSize _ m i - k = proj₁ reasoning - i+k≡n = proj₂ (proj₂ reasoning) - eq = proj₁ (proj₂ reasoning) - low = take′ t (Fin.toℕ i) (cast′ t eq x) - middle = take′ t m (drop′ t (Fin.toℕ i) (cast′ t eq x)) - high = drop′ t m (drop′ t (Fin.toℕ i) (cast′ t eq x)) - -cut : ∀ t → Term Σ Γ Δ (asType t (n ℕ.+ m)) → Term Σ Γ Δ (fin (suc n)) → Term Σ Γ Δ (tuple _ (asType t m ∷ asType t n ∷ [])) -cut τ t t′ = func₂ (cut′ τ) t t′ - -flatten : ∀ {ms : Vec ℕ n} → ⟦ Vec.map fin ms ⟧ₜ′ → All Fin ms -flatten {ms = []} _ = [] -flatten {ms = _ ∷ ms} (lift x , xs) = x ∷ flatten xs - -𝒦 : Literal t → Term Σ Γ Δ t -𝒦 (x ′b) = func₀ (lift x) -𝒦 (x ′i) = func₀ (lift (x ℤ′.×′ 1ℤ)) -𝒦 (x ′f) = func₀ (lift x) -𝒦 (x ′r) = func₀ (lift (x ℝ′.×′ 1ℝ)) -𝒦 (x ′x) = func₀ (lift (Bool.if x then 1𝔹 else 0𝔹)) -𝒦 ([] ′xs) = func₀ [] -𝒦 ((x ∷ xs) ′xs) = func₂ (flip Vec._∷ʳ_) (𝒦 (x ′x)) (𝒦 (xs ′xs)) -𝒦 (x ′a) = func₁ Vec.replicate (𝒦 x) - -module _ (2≉0 : 2≉0) where - ℰ : Code.Expression Σ Γ t → Term Σ Γ Δ t - ℰ e = (uncurry helper) (M.elimFunctions e) - where - 1+m⊔n≡1+k : ∀ m n → ∃ λ k → suc m ℕ.⊔ n ≡ suc k - 1+m⊔n≡1+k m 0 = m , refl - 1+m⊔n≡1+k m (suc n) = m ℕ.⊔ n , refl - - pull-0₂ : ∀ {x y} → x ℕ.⊔ y ≡ 0 → x ≡ 0 - pull-0₂ {0} {0} refl = refl - pull-0₂ {0} {suc y} () - - pull-0₃ : ∀ {x y z} → x ℕ.⊔ y ℕ.⊔ z ≡ 0 → x ≡ 0 - pull-0₃ {0} {0} {0} refl = refl - pull-0₃ {0} {suc y} {0} () - pull-0₃ {suc x} {0} {0} () - pull-0₃ {suc x} {0} {suc z} () - - pull-1₃ : ∀ x {y z} → x ℕ.⊔ y ℕ.⊔ z ≡ 0 → y ≡ 0 - pull-1₃ 0 {0} {0} refl = refl - pull-1₃ 0 {suc y} {0} () - pull-1₃ (suc x) {0} {0} () - pull-1₃ (suc x) {0} {suc z} () - - pull-last : ∀ {x y} → x ℕ.⊔ y ≡ 0 → y ≡ 0 - pull-last {0} {0} refl = refl - pull-last {suc x} {0} () - - helper : ∀ (e : Code.Expression Σ Γ t) → M.callDepth e ≡ 0 → Term Σ Γ Δ t - helper (Code.lit x) eq = 𝒦 x - helper (Code.state i) eq = state i - helper (Code.var i) eq = var i - helper (Code.abort e) eq = func₁ (λ ()) (helper e eq) - helper (Code._≟_ {hasEquality = hasEq} e e₁) eq = [ toWitness hasEq ][ helper e (pull-0₂ eq) ≟ helper e₁ (pull-last eq) ] - helper (Code._<?_ {isNumeric = isNum} e e₁) eq = [ toWitness isNum ][ helper e (pull-0₂ eq) <? helper e₁ (pull-last eq) ] - helper (Code.inv e) eq = func₁ (lift ∘ Bool.not ∘ lower) (helper e eq) - helper (e Code.&& e₁) eq = func₂ (λ (lift x) (lift y) → lift (x Bool.∧ y)) (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (e Code.|| e₁) eq = func₂ (λ (lift x) (lift y) → lift (x Bool.∨ y)) (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (Code.not e) eq = func₁ (Vec.map (lift ∘ 𝔹.¬_ ∘ lower)) (helper e eq) - helper (e Code.and e₁) eq = func₂ (λ xs ys → Vec.zipWith (λ (lift x) (lift y) → lift (x 𝔹.∧ y)) xs ys) (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (e Code.or e₁) eq = func₂ (λ xs ys → Vec.zipWith (λ (lift x) (lift y) → lift (x 𝔹.∨ y)) xs ys) (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (Code.[_] {t = t} e) eq = [ t ][ helper e eq ] - helper (Code.unbox {t = t} e) eq = unbox t (helper e eq) - helper (Code.splice {t = t} e e₁ e₂) eq = splice t (helper e (pull-0₃ eq)) (helper e₁ (pull-1₃ (M.callDepth e) eq)) (helper e₂ (pull-last eq)) - helper (Code.cut {t = t} e e₁) eq = cut t (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (Code.cast {t = t} i≡j e) eq = cast t i≡j (helper e eq) - helper (Code.-_ {isNumeric = isNum} e) eq = [ toWitness isNum ][- helper e eq ] - helper (Code._+_ {isNumeric₁ = isNum₁} {isNumeric₂ = isNum₂} e e₁) eq = [ _ , _ , isNum₁ , isNum₂ ][ helper e (pull-0₂ eq) + helper e₁ (pull-last eq) ] - helper (Code._*_ {isNumeric₁ = isNum₁} {isNumeric₂ = isNum₂} e e₁) eq = [ _ , _ , isNum₁ , isNum₂ ][ helper e (pull-0₂ eq) * helper e₁ (pull-last eq) ] - helper (Code._^_ {isNumeric = isNum} e y) eq = [ toWitness isNum ][ helper e eq ^ y ] - helper (e Code.>> n) eq = [ 2≉0 ][ helper e eq >> n ] - helper (Code.rnd e) eq = func₁ (lift ∘ ⌊_⌋ ∘ lower) (helper e eq) - helper (Code.fin f e) eq = func₁ (lift ∘ f ∘ flatten) (helper e eq) - helper (Code.asInt e) eq = func₁ (lift ∘ (ℤ′._×′ 1ℤ) ∘ Fin.toℕ ∘ lower) (helper e eq) - helper Code.nil eq = func₀ _ - helper (Code.cons e e₁) eq = func₂ _,_ (helper e (pull-0₂ eq)) (helper e₁ (pull-last eq)) - helper (Code.head e) eq = func₁ proj₁ (helper e eq) - helper (Code.tail e) eq = func₁ proj₂ (helper e eq) - helper (Code.call f es) eq = contradiction (trans (sym eq) (proj₂ (1+m⊔n≡1+k (M.funCallDepth f) (M.callDepth′ es)))) ℕₚ.0≢1+n - helper (Code.if e then e₁ else e₂) eq = func (λ (lift b , x , x₁ , _) → Bool.if b then x else x₁) (helper e (pull-0₃ eq) ∷ helper e₁ (pull-1₃ (M.callDepth e) eq) ∷ helper e₂ (pull-last eq) ∷ []) - - ℰ′ : All (Code.Expression Σ Γ) ts → All (Term Σ Γ Δ) ts - ℰ′ [] = [] - ℰ′ (e ∷ es) = ℰ e ∷ ℰ′ es - - subst : Term Σ Γ Δ t → {e : Code.Expression Σ Γ t′} → Code.CanAssign Σ e → Term Σ Γ Δ t′ → Term Σ Γ Δ t - subst t (Code.state i) t′ = substState t i t′ - subst t (Code.var i) t′ = substVar t i t′ - subst t (Code.abort e) t′ = func₁ (λ ()) (ℰ e) - subst t (Code.[_] {t = τ} ref) t′ = subst t ref (unbox τ t′) - subst t (Code.unbox {t = τ} ref) t′ = subst t ref [ τ ][ t′ ] - subst t (Code.splice {t = τ} ref ref₁ e₃) t′ = subst (subst t ref (func₁ proj₁ (cut τ t′ (ℰ e₃)))) ref₁ (func₁ (proj₁ ∘ proj₂) (cut τ t′ (ℰ e₃))) - subst t (Code.cut {t = τ} ref e₂) t′ = subst t ref (splice τ (func₁ proj₁ t′) (func₁ (proj₁ ∘ proj₂) t′) (ℰ e₂)) - subst t (Code.cast {t = τ} eq ref) t′ = subst t ref (cast τ (sym eq) t′) - subst t Code.nil t′ = t - subst t (Code.cons ref ref₁) t′ = subst (subst t ref (func₁ proj₁ t′)) ref₁ (func₁ proj₂ t′) - subst t (Code.head {e = e} ref) t′ = subst t ref (func₂ _,_ t′ (func₁ proj₂ (ℰ e))) - subst t (Code.tail {t = τ} {e = e} ref) t′ = subst t ref (func₂ {t₁ = τ} _,_ (func₁ proj₁ (ℰ e)) t′) - -⟦_⟧ : Term Σ Γ Δ t → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → ⟦ t ⟧ₜ -⟦_⟧′ : ∀ {ts} → All (Term Σ Γ Δ) {n} ts → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → ⟦ ts ⟧ₜ′ - -⟦ state i ⟧ σ γ δ = fetch i σ -⟦ var i ⟧ σ γ δ = fetch i γ -⟦ meta i ⟧ σ γ δ = fetch i δ -⟦ func f ts ⟧ σ γ δ = f (⟦ ts ⟧′ σ γ δ) - -⟦ [] ⟧′ σ γ δ = _ -⟦ t ∷ ts ⟧′ σ γ δ = ⟦ t ⟧ σ γ δ , ⟦ ts ⟧′ σ γ δ + ⊔-inj : ∀ i xs → ⨆[ n ] xs ≡ 0 → Vecᵣ.lookup i xs ≡ 0 + ⊔-inj i xs eq = ℕₚ.n≤0⇒n≡0 (ℕₚ.≤-trans (lookup-⨆-≤ i xs) (ℕₚ.≤-reflexive eq)) + + go : ∀ (e : Expression Σ Γ t) → CallDepth.expr e ≡ 0 → Term Σ Γ Δ t + go (lit {t} x) ≡0 = lit (Κ[ t ] x) + go (state i) ≡0 = state i + go (var i) ≡0 = var i + go (e ≟ e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) ≟ go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (e <? e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) <? go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (inv e) ≡0 = inv (go e ≡0) + go (e && e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) && go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (e || e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) || go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (not e) ≡0 = not (go e ≡0) + go (e and e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) and go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (e or e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) or go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go [ e ] ≡0 = [ go e ≡0 ] + go (unbox e) ≡0 = unbox (go e ≡0) + go (merge e e₁ e₂) ≡0 = merge (go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0)) (go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0)) (go e₂ (⊔-inj 2F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0)) + go (slice e e₁) ≡0 = slice (go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) (go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) + go (cut e e₁) ≡0 = cut (go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) (go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) + go (cast eq e) ≡0 = cast eq (go e ≡0) + go (- e) ≡0 = - go e ≡0 + go (e + e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (e * e₁) ≡0 = go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0) * go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0) + go (e ^ x) ≡0 = go e ≡0 ^ x + go (e >> n) ≡0 = go e ≡0 >> n + go (rnd e) ≡0 = rnd (go e ≡0) + go (fin f e) ≡0 = fin f (go e ≡0) + go (asInt e) ≡0 = asInt (go e ≡0) + go nil ≡0 = nil + go (cons e e₁) ≡0 = cons (go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) (go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁) ≡0)) + go (head e) ≡0 = head (go e ≡0) + go (tail e) ≡0 = tail (go e ≡0) + go (call f es) ≡0 = ⊥-elim (ℕₚ.>⇒≢ (CallDepth.call>0 f es) ≡0) + go (if e then e₁ else e₂) ≡0 = if go e (⊔-inj 0F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0) then go e₁ (⊔-inj 1F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0) else go e₂ (⊔-inj 2F (CallDepth.expr e , CallDepth.expr e₁ , CallDepth.expr e₂) ≡0) + +module Cast where + type : t ≡ t′ → Term Σ Γ Δ t → Term Σ Γ Δ t′ + type refl = id + +module State where + subst : ∀ i → Term Σ Γ Δ t → Term Σ Γ Δ (lookup Σ i) → Term Σ Γ Δ t + subst i (lit x) e′ = lit x + subst i (state j) e′ with i Fin.≟ j + ... | yes refl = e′ + ... | no i≢j = state j + subst i (var j) e′ = var j + subst i (meta j) e′ = meta j + subst i (e ≟ e₁) e′ = subst i e e′ ≟ subst i e₁ e′ + subst i (e <? e₁) e′ = subst i e e′ <? subst i e₁ e′ + subst i (inv e) e′ = inv (subst i e e′) + subst i (e && e₁) e′ = subst i e e′ && subst i e₁ e′ + subst i (e || e₁) e′ = subst i e e′ || subst i e₁ e′ + subst i (not e) e′ = not (subst i e e′) + subst i (e and e₁) e′ = subst i e e′ and subst i e₁ e′ + subst i (e or e₁) e′ = subst i e e′ or subst i e₁ e′ + subst i [ e ] e′ = [ subst i e e′ ] + subst i (unbox e) e′ = unbox (subst i e e′) + subst i (merge e e₁ e₂) e′ = merge (subst i e e′) (subst i e₁ e′) (subst i e₂ e′) + subst i (slice e e₁) e′ = slice (subst i e e′) (subst i e₁ e′) + subst i (cut e e₁) e′ = cut (subst i e e′) (subst i e₁ e′) + subst i (cast eq e) e′ = cast eq (subst i e e′) + subst i (- e) e′ = - subst i e e′ + subst i (e + e₁) e′ = subst i e e′ + subst i e₁ e′ + subst i (e * e₁) e′ = subst i e e′ * subst i e₁ e′ + subst i (e ^ x) e′ = subst i e e′ ^ x + subst i (e >> n) e′ = subst i e e′ >> n + subst i (rnd e) e′ = rnd (subst i e e′) + subst i (fin f e) e′ = fin f (subst i e e′) + subst i (asInt e) e′ = asInt (subst i e e′) + subst i nil e′ = nil + subst i (cons e e₁) e′ = cons (subst i e e′) (subst i e₁ e′) + subst i (head e) e′ = head (subst i e e′) + subst i (tail e) e′ = tail (subst i e e′) + subst i (if e then e₁ else e₂) e′ = if subst i e e′ then subst i e₁ e′ else subst i e₂ e′ + +module Var {Γ : Vec Type o} where + weaken : ∀ i → Term Σ Γ Δ t → Term Σ (insert Γ i t′) Δ t + weaken i (lit x) = lit x + weaken i (state j) = state j + weaken i (var j) = Cast.type (Vecₚ.insert-punchIn _ i _ j) (var (Fin.punchIn i j)) + weaken i (meta j) = meta j + weaken i (e ≟ e₁) = weaken i e ≟ weaken i e₁ + weaken i (e <? e₁) = weaken i e <? weaken i e₁ + weaken i (inv e) = inv (weaken i e) + weaken i (e && e₁) = weaken i e && weaken i e₁ + weaken i (e || e₁) = weaken i e || weaken i e₁ + weaken i (not e) = not (weaken i e) + weaken i (e and e₁) = weaken i e and weaken i e₁ + weaken i (e or e₁) = weaken i e or weaken i e₁ + weaken i [ e ] = [ weaken i e ] + weaken i (unbox e) = unbox (weaken i e) + weaken i (merge e e₁ e₂) = merge (weaken i e) (weaken i e₁) (weaken i e₂) + weaken i (slice e e₁) = slice (weaken i e) (weaken i e₁) + weaken i (cut e e₁) = cut (weaken i e) (weaken i e₁) + weaken i (cast eq e) = cast eq (weaken i e) + weaken i (- e) = - weaken i e + weaken i (e + e₁) = weaken i e + weaken i e₁ + weaken i (e * e₁) = weaken i e * weaken i e₁ + weaken i (e ^ x) = weaken i e ^ x + weaken i (e >> n) = weaken i e >> n + weaken i (rnd e) = rnd (weaken i e) + weaken i (fin f e) = fin f (weaken i e) + weaken i (asInt e) = asInt (weaken i e) + weaken i nil = nil + weaken i (cons e e₁) = cons (weaken i e) (weaken i e₁) + weaken i (head e) = head (weaken i e) + weaken i (tail e) = tail (weaken i e) + weaken i (if e then e₁ else e₂) = if weaken i e then weaken i e₁ else weaken i e₂ + + weakenAll : Term Σ [] Δ t → Term Σ Γ Δ t + weakenAll (lit x) = lit x + weakenAll (state j) = state j + weakenAll (meta j) = meta j + weakenAll (e ≟ e₁) = weakenAll e ≟ weakenAll e₁ + weakenAll (e <? e₁) = weakenAll e <? weakenAll e₁ + weakenAll (inv e) = inv (weakenAll e) + weakenAll (e && e₁) = weakenAll e && weakenAll e₁ + weakenAll (e || e₁) = weakenAll e || weakenAll e₁ + weakenAll (not e) = not (weakenAll e) + weakenAll (e and e₁) = weakenAll e and weakenAll e₁ + weakenAll (e or e₁) = weakenAll e or weakenAll e₁ + weakenAll [ e ] = [ weakenAll e ] + weakenAll (unbox e) = unbox (weakenAll e) + weakenAll (merge e e₁ e₂) = merge (weakenAll e) (weakenAll e₁) (weakenAll e₂) + weakenAll (slice e e₁) = slice (weakenAll e) (weakenAll e₁) + weakenAll (cut e e₁) = cut (weakenAll e) (weakenAll e₁) + weakenAll (cast eq e) = cast eq (weakenAll e) + weakenAll (- e) = - weakenAll e + weakenAll (e + e₁) = weakenAll e + weakenAll e₁ + weakenAll (e * e₁) = weakenAll e * weakenAll e₁ + weakenAll (e ^ x) = weakenAll e ^ x + weakenAll (e >> n) = weakenAll e >> n + weakenAll (rnd e) = rnd (weakenAll e) + weakenAll (fin f e) = fin f (weakenAll e) + weakenAll (asInt e) = asInt (weakenAll e) + weakenAll nil = nil + weakenAll (cons e e₁) = cons (weakenAll e) (weakenAll e₁) + weakenAll (head e) = head (weakenAll e) + weakenAll (tail e) = tail (weakenAll e) + weakenAll (if e then e₁ else e₂) = if weakenAll e then weakenAll e₁ else weakenAll e₂ + + inject : ∀ (ts : Vec Type n) → Term Σ Γ Δ t → Term Σ (Γ ++ ts) Δ t + inject ts (lit x) = lit x + inject ts (state j) = state j + inject ts (var j) = Cast.type (Vecₚ.lookup-++ˡ Γ ts j) (var (Fin.inject+ _ j)) + inject ts (meta j) = meta j + inject ts (e ≟ e₁) = inject ts e ≟ inject ts e₁ + inject ts (e <? e₁) = inject ts e <? inject ts e₁ + inject ts (inv e) = inv (inject ts e) + inject ts (e && e₁) = inject ts e && inject ts e₁ + inject ts (e || e₁) = inject ts e || inject ts e₁ + inject ts (not e) = not (inject ts e) + inject ts (e and e₁) = inject ts e and inject ts e₁ + inject ts (e or e₁) = inject ts e or inject ts e₁ + inject ts [ e ] = [ inject ts e ] + inject ts (unbox e) = unbox (inject ts e) + inject ts (merge e e₁ e₂) = merge (inject ts e) (inject ts e₁) (inject ts e₂) + inject ts (slice e e₁) = slice (inject ts e) (inject ts e₁) + inject ts (cut e e₁) = cut (inject ts e) (inject ts e₁) + inject ts (cast eq e) = cast eq (inject ts e) + inject ts (- e) = - inject ts e + inject ts (e + e₁) = inject ts e + inject ts e₁ + inject ts (e * e₁) = inject ts e * inject ts e₁ + inject ts (e ^ x) = inject ts e ^ x + inject ts (e >> n) = inject ts e >> n + inject ts (rnd e) = rnd (inject ts e) + inject ts (fin f e) = fin f (inject ts e) + inject ts (asInt e) = asInt (inject ts e) + inject ts nil = nil + inject ts (cons e e₁) = cons (inject ts e) (inject ts e₁) + inject ts (head e) = head (inject ts e) + inject ts (tail e) = tail (inject ts e) + inject ts (if e then e₁ else e₂) = if inject ts e then inject ts e₁ else inject ts e₂ + + raise : ∀ (ts : Vec Type n) → Term Σ Γ Δ t → Term Σ (ts ++ Γ) Δ t + raise ts (lit x) = lit x + raise ts (state j) = state j + raise ts (var j) = Cast.type (Vecₚ.lookup-++ʳ ts Γ j) (var (Fin.raise _ j)) + raise ts (meta j) = meta j + raise ts (e ≟ e₁) = raise ts e ≟ raise ts e₁ + raise ts (e <? e₁) = raise ts e <? raise ts e₁ + raise ts (inv e) = inv (raise ts e) + raise ts (e && e₁) = raise ts e && raise ts e₁ + raise ts (e || e₁) = raise ts e || raise ts e₁ + raise ts (not e) = not (raise ts e) + raise ts (e and e₁) = raise ts e and raise ts e₁ + raise ts (e or e₁) = raise ts e or raise ts e₁ + raise ts [ e ] = [ raise ts e ] + raise ts (unbox e) = unbox (raise ts e) + raise ts (merge e e₁ e₂) = merge (raise ts e) (raise ts e₁) (raise ts e₂) + raise ts (slice e e₁) = slice (raise ts e) (raise ts e₁) + raise ts (cut e e₁) = cut (raise ts e) (raise ts e₁) + raise ts (cast eq e) = cast eq (raise ts e) + raise ts (- e) = - raise ts e + raise ts (e + e₁) = raise ts e + raise ts e₁ + raise ts (e * e₁) = raise ts e * raise ts e₁ + raise ts (e ^ x) = raise ts e ^ x + raise ts (e >> n) = raise ts e >> n + raise ts (rnd e) = rnd (raise ts e) + raise ts (fin f e) = fin f (raise ts e) + raise ts (asInt e) = asInt (raise ts e) + raise ts nil = nil + raise ts (cons e e₁) = cons (raise ts e) (raise ts e₁) + raise ts (head e) = head (raise ts e) + raise ts (tail e) = tail (raise ts e) + raise ts (if e then e₁ else e₂) = if raise ts e then raise ts e₁ else raise ts e₂ + + elim : ∀ i → Term Σ (insert Γ i t′) Δ t → Term Σ Γ Δ t′ → Term Σ Γ Δ t + elim i (lit x) e′ = lit x + elim i (state j) e′ = state j + elim i (var j) e′ with i Fin.≟ j + ... | yes refl = Cast.type (sym (Vecₚ.insert-lookup Γ i _)) e′ + ... | no i≢j = Cast.type (punchOut-insert Γ i≢j _) (var (Fin.punchOut i≢j)) + elim i (meta j) e′ = meta j + elim i (e ≟ e₁) e′ = elim i e e′ ≟ elim i e₁ e′ + elim i (e <? e₁) e′ = elim i e e′ <? elim i e₁ e′ + elim i (inv e) e′ = inv (elim i e e′) + elim i (e && e₁) e′ = elim i e e′ && elim i e₁ e′ + elim i (e || e₁) e′ = elim i e e′ || elim i e₁ e′ + elim i (not e) e′ = not (elim i e e′) + elim i (e and e₁) e′ = elim i e e′ and elim i e₁ e′ + elim i (e or e₁) e′ = elim i e e′ or elim i e₁ e′ + elim i [ e ] e′ = [ elim i e e′ ] + elim i (unbox e) e′ = unbox (elim i e e′) + elim i (merge e e₁ e₂) e′ = merge (elim i e e′) (elim i e₁ e′) (elim i e₂ e′) + elim i (slice e e₁) e′ = slice (elim i e e′) (elim i e₁ e′) + elim i (cut e e₁) e′ = cut (elim i e e′) (elim i e₁ e′) + elim i (cast eq e) e′ = cast eq (elim i e e′) + elim i (- e) e′ = - elim i e e′ + elim i (e + e₁) e′ = elim i e e′ + elim i e₁ e′ + elim i (e * e₁) e′ = elim i e e′ * elim i e₁ e′ + elim i (e ^ x) e′ = elim i e e′ ^ x + elim i (e >> n) e′ = elim i e e′ >> n + elim i (rnd e) e′ = rnd (elim i e e′) + elim i (fin f e) e′ = fin f (elim i e e′) + elim i (asInt e) e′ = asInt (elim i e e′) + elim i nil e′ = nil + elim i (cons e e₁) e′ = cons (elim i e e′) (elim i e₁ e′) + elim i (head e) e′ = head (elim i e e′) + elim i (tail e) e′ = tail (elim i e e′) + elim i (if e then e₁ else e₂) e′ = if elim i e e′ then elim i e₁ e′ else elim i e₂ e′ + + elimAll : Term Σ Γ Δ t → All (Term Σ ts Δ) Γ → Term Σ ts Δ t + elimAll (lit x) es = lit x + elimAll (state j) es = state j + elimAll (var j) es = All.lookup j es + elimAll (meta j) es = meta j + elimAll (e ≟ e₁) es = elimAll e es ≟ elimAll e₁ es + elimAll (e <? e₁) es = elimAll e es <? elimAll e₁ es + elimAll (inv e) es = inv (elimAll e es) + elimAll (e && e₁) es = elimAll e es && elimAll e₁ es + elimAll (e || e₁) es = elimAll e es || elimAll e₁ es + elimAll (not e) es = not (elimAll e es) + elimAll (e and e₁) es = elimAll e es and elimAll e₁ es + elimAll (e or e₁) es = elimAll e es or elimAll e₁ es + elimAll [ e ] es = [ elimAll e es ] + elimAll (unbox e) es = unbox (elimAll e es) + elimAll (merge e e₁ e₂) es = merge (elimAll e es) (elimAll e₁ es) (elimAll e₂ es) + elimAll (slice e e₁) es = slice (elimAll e es) (elimAll e₁ es) + elimAll (cut e e₁) es = cut (elimAll e es) (elimAll e₁ es) + elimAll (cast eq e) es = cast eq (elimAll e es) + elimAll (- e) es = - elimAll e es + elimAll (e + e₁) es = elimAll e es + elimAll e₁ es + elimAll (e * e₁) es = elimAll e es * elimAll e₁ es + elimAll (e ^ x) es = elimAll e es ^ x + elimAll (e >> n) es = elimAll e es >> n + elimAll (rnd e) es = rnd (elimAll e es) + elimAll (fin f e) es = fin f (elimAll e es) + elimAll (asInt e) es = asInt (elimAll e es) + elimAll nil es = nil + elimAll (cons e e₁) es = cons (elimAll e es) (elimAll e₁ es) + elimAll (head e) es = head (elimAll e es) + elimAll (tail e) es = tail (elimAll e es) + elimAll (if e then e₁ else e₂) es = if elimAll e es then elimAll e₁ es else elimAll e₂ es + + subst : ∀ i → Term Σ Γ Δ t → Term Σ Γ Δ (lookup Γ i) → Term Σ Γ Δ t + subst i (lit x) e′ = lit x + subst i (state j) e′ = state j + subst i (var j) e′ with i Fin.≟ j + ... | yes refl = e′ + ... | no i≢j = var j + subst i (meta j) e′ = meta j + subst i (e ≟ e₁) e′ = subst i e e′ ≟ subst i e₁ e′ + subst i (e <? e₁) e′ = subst i e e′ <? subst i e₁ e′ + subst i (inv e) e′ = inv (subst i e e′) + subst i (e && e₁) e′ = subst i e e′ && subst i e₁ e′ + subst i (e || e₁) e′ = subst i e e′ || subst i e₁ e′ + subst i (not e) e′ = not (subst i e e′) + subst i (e and e₁) e′ = subst i e e′ and subst i e₁ e′ + subst i (e or e₁) e′ = subst i e e′ or subst i e₁ e′ + subst i [ e ] e′ = [ subst i e e′ ] + subst i (unbox e) e′ = unbox (subst i e e′) + subst i (merge e e₁ e₂) e′ = merge (subst i e e′) (subst i e₁ e′) (subst i e₂ e′) + subst i (slice e e₁) e′ = slice (subst i e e′) (subst i e₁ e′) + subst i (cut e e₁) e′ = cut (subst i e e′) (subst i e₁ e′) + subst i (cast eq e) e′ = cast eq (subst i e e′) + subst i (- e) e′ = - subst i e e′ + subst i (e + e₁) e′ = subst i e e′ + subst i e₁ e′ + subst i (e * e₁) e′ = subst i e e′ * subst i e₁ e′ + subst i (e ^ x) e′ = subst i e e′ ^ x + subst i (e >> n) e′ = subst i e e′ >> n + subst i (rnd e) e′ = rnd (subst i e e′) + subst i (fin f e) e′ = fin f (subst i e e′) + subst i (asInt e) e′ = asInt (subst i e e′) + subst i nil e′ = nil + subst i (cons e e₁) e′ = cons (subst i e e′) (subst i e₁ e′) + subst i (head e) e′ = head (subst i e e′) + subst i (tail e) e′ = tail (subst i e e′) + subst i (if e then e₁ else e₂) e′ = if subst i e e′ then subst i e₁ e′ else subst i e₂ e′ + +module Meta {Δ : Vec Type o} where + weaken : ∀ i → Term Σ Γ Δ t → Term Σ Γ (insert Δ i t′) t + weaken i (lit x) = lit x + weaken i (state j) = state j + weaken i (var j) = var j + weaken i (meta j) = Cast.type (Vecₚ.insert-punchIn _ i _ j) (meta (Fin.punchIn i j)) + weaken i (e ≟ e₁) = weaken i e ≟ weaken i e₁ + weaken i (e <? e₁) = weaken i e <? weaken i e₁ + weaken i (inv e) = inv (weaken i e) + weaken i (e && e₁) = weaken i e && weaken i e₁ + weaken i (e || e₁) = weaken i e || weaken i e₁ + weaken i (not e) = not (weaken i e) + weaken i (e and e₁) = weaken i e and weaken i e₁ + weaken i (e or e₁) = weaken i e or weaken i e₁ + weaken i [ e ] = [ weaken i e ] + weaken i (unbox e) = unbox (weaken i e) + weaken i (merge e e₁ e₂) = merge (weaken i e) (weaken i e₁) (weaken i e₂) + weaken i (slice e e₁) = slice (weaken i e) (weaken i e₁) + weaken i (cut e e₁) = cut (weaken i e) (weaken i e₁) + weaken i (cast eq e) = cast eq (weaken i e) + weaken i (- e) = - weaken i e + weaken i (e + e₁) = weaken i e + weaken i e₁ + weaken i (e * e₁) = weaken i e * weaken i e₁ + weaken i (e ^ x) = weaken i e ^ x + weaken i (e >> n) = weaken i e >> n + weaken i (rnd e) = rnd (weaken i e) + weaken i (fin f e) = fin f (weaken i e) + weaken i (asInt e) = asInt (weaken i e) + weaken i nil = nil + weaken i (cons e e₁) = cons (weaken i e) (weaken i e₁) + weaken i (head e) = head (weaken i e) + weaken i (tail e) = tail (weaken i e) + weaken i (if e then e₁ else e₂) = if weaken i e then weaken i e₁ else weaken i e₂ + + elim : ∀ i → Term Σ Γ (insert Δ i t′) t → Term Σ Γ Δ t′ → Term Σ Γ Δ t + elim i (lit x) e′ = lit x + elim i (state j) e′ = state j + elim i (var j) e′ = var j + elim i (meta j) e′ with i Fin.≟ j + ... | yes refl = Cast.type (sym (Vecₚ.insert-lookup Δ i _)) e′ + ... | no i≢j = Cast.type (punchOut-insert Δ i≢j _) (meta (Fin.punchOut i≢j)) + elim i (e ≟ e₁) e′ = elim i e e′ ≟ elim i e₁ e′ + elim i (e <? e₁) e′ = elim i e e′ <? elim i e₁ e′ + elim i (inv e) e′ = inv (elim i e e′) + elim i (e && e₁) e′ = elim i e e′ && elim i e₁ e′ + elim i (e || e₁) e′ = elim i e e′ || elim i e₁ e′ + elim i (not e) e′ = not (elim i e e′) + elim i (e and e₁) e′ = elim i e e′ and elim i e₁ e′ + elim i (e or e₁) e′ = elim i e e′ or elim i e₁ e′ + elim i [ e ] e′ = [ elim i e e′ ] + elim i (unbox e) e′ = unbox (elim i e e′) + elim i (merge e e₁ e₂) e′ = merge (elim i e e′) (elim i e₁ e′) (elim i e₂ e′) + elim i (slice e e₁) e′ = slice (elim i e e′) (elim i e₁ e′) + elim i (cut e e₁) e′ = cut (elim i e e′) (elim i e₁ e′) + elim i (cast eq e) e′ = cast eq (elim i e e′) + elim i (- e) e′ = - elim i e e′ + elim i (e + e₁) e′ = elim i e e′ + elim i e₁ e′ + elim i (e * e₁) e′ = elim i e e′ * elim i e₁ e′ + elim i (e ^ x) e′ = elim i e e′ ^ x + elim i (e >> n) e′ = elim i e e′ >> n + elim i (rnd e) e′ = rnd (elim i e e′) + elim i (fin f e) e′ = fin f (elim i e e′) + elim i (asInt e) e′ = asInt (elim i e e′) + elim i nil e′ = nil + elim i (cons e e₁) e′ = cons (elim i e e′) (elim i e₁ e′) + elim i (head e) e′ = head (elim i e e′) + elim i (tail e) e′ = tail (elim i e e′) + elim i (if e then e₁ else e₂) e′ = if elim i e e′ then elim i e₁ e′ else elim i e₂ e′ + +subst : Term Σ Γ Δ t → Reference Σ Γ t′ → Term Σ Γ Δ t′ → Term Σ Γ Δ t +subst e (state i) val = State.subst i e val +subst e (var i) val = Var.subst i e val +subst e [ ref ] val = subst e ref (unbox val) +subst e (unbox ref) val = subst e ref [ val ] +subst e (merge ref ref₁ x) val = subst (subst e ref (slice val (↓ x))) ref₁ (cut val (↓ x)) +subst e (slice ref x) val = subst e ref (merge val (↓ ! cut ref x) (↓ x)) +subst e (cut ref x) val = subst e ref (merge (↓ ! slice ref x) val (↓ x)) +subst e (cast eq ref) val = subst e ref (cast (sym eq) val) +subst e nil val = e +subst e (cons ref ref₁) val = subst (subst e ref (head val)) ref₁ (tail val) +subst e (head ref) val = subst e ref (cons val (↓ ! tail ref)) +subst e (tail ref) val = subst e ref (cons (↓ ! head ref) val) + +module Semantics (2≉0 : 2≉0) {Σ : Vec Type i} {Γ : Vec Type j} {Δ : Vec Type k} where + ⟦_⟧ : Term Σ Γ Δ t → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Δ ⟧ₜ′ → ⟦ t ⟧ₜ + ⟦ lit x ⟧ σ γ δ = x + ⟦ state i ⟧ σ γ δ = fetch i Σ σ + ⟦ var i ⟧ σ γ δ = fetch i Γ γ + ⟦ meta i ⟧ σ γ δ = fetch i Δ δ + ⟦ e ≟ e₁ ⟧ σ γ δ = (lift ∘₂ does ∘₂ ≈-dec) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ e <? e₁ ⟧ σ γ δ = (lift ∘₂ does ∘₂ <-dec) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ inv e ⟧ σ γ δ = lift ∘ Bool.not ∘ lower $ ⟦ e ⟧ σ γ δ + ⟦ e && e₁ ⟧ σ γ δ = (lift ∘₂ Bool._∧_ on lower) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ e || e₁ ⟧ σ γ δ = (lift ∘₂ Bool._∨_ on lower) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ not e ⟧ σ γ δ = map (lift ∘ 𝔹.¬_ ∘ lower) (⟦ e ⟧ σ γ δ) + ⟦ e and e₁ ⟧ σ γ δ = zipWith (lift ∘₂ 𝔹._∧_ on lower) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ e or e₁ ⟧ σ γ δ = zipWith (lift ∘₂ 𝔹._∨_ on lower) (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ [ e ] ⟧ σ γ δ = ⟦ e ⟧ σ γ δ ∷ [] + ⟦ unbox e ⟧ σ γ δ = Vec.head (⟦ e ⟧ σ γ δ) + ⟦ merge e e₁ e₂ ⟧ σ γ δ = mergeVec (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) (lower (⟦ e₂ ⟧ σ γ δ)) + ⟦ slice e e₁ ⟧ σ γ δ = sliceVec (⟦ e ⟧ σ γ δ) (lower (⟦ e₁ ⟧ σ γ δ)) + ⟦ cut e e₁ ⟧ σ γ δ = cutVec (⟦ e ⟧ σ γ δ) (lower (⟦ e₁ ⟧ σ γ δ)) + ⟦ cast eq e ⟧ σ γ δ = castVec eq (⟦ e ⟧ σ γ δ) + ⟦ - e ⟧ σ γ δ = neg (⟦ e ⟧ σ γ δ) + ⟦ e + e₁ ⟧ σ γ δ = add (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ e * e₁ ⟧ σ γ δ = mul (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ e ^ x ⟧ σ γ δ = pow (⟦ e ⟧ σ γ δ) x + ⟦ e >> n ⟧ σ γ δ = lift ∘ flip (shift 2≉0) n ∘ lower $ ⟦ e ⟧ σ γ δ + ⟦ rnd e ⟧ σ γ δ = lift ∘ ⌊_⌋ ∘ lower $ ⟦ e ⟧ σ γ δ + ⟦ fin {ms = ms} f e ⟧ σ γ δ = lift ∘ f ∘ lowerFin ms $ ⟦ e ⟧ σ γ δ + ⟦ asInt e ⟧ σ γ δ = lift ∘ 𝕀⇒ℤ ∘ 𝕀.+_ ∘ Fin.toℕ ∘ lower $ ⟦ e ⟧ σ γ δ + ⟦ nil ⟧ σ γ δ = _ + ⟦ cons {ts = ts} e e₁ ⟧ σ γ δ = cons′ ts (⟦ e ⟧ σ γ δ) (⟦ e₁ ⟧ σ γ δ) + ⟦ head {ts = ts} e ⟧ σ γ δ = head′ ts (⟦ e ⟧ σ γ δ) + ⟦ tail {ts = ts} e ⟧ σ γ δ = tail′ ts (⟦ e ⟧ σ γ δ) + ⟦ if e then e₁ else e₂ ⟧ σ γ δ = Bool.if lower (⟦ e ⟧ σ γ δ) then ⟦ e₁ ⟧ σ γ δ else ⟦ e₂ ⟧ σ γ δ diff --git a/src/Helium/Semantics/Axiomatic/Triple.agda b/src/Helium/Semantics/Axiomatic/Triple.agda index 23a487e..8c6b45a 100644 --- a/src/Helium/Semantics/Axiomatic/Triple.agda +++ b/src/Helium/Semantics/Axiomatic/Triple.agda @@ -7,40 +7,61 @@ {-# OPTIONS --safe --without-K #-} open import Helium.Data.Pseudocode.Algebra using (RawPseudocode) +import Helium.Semantics.Core as Core module Helium.Semantics.Axiomatic.Triple {b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃} (rawPseudocode : RawPseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) + (2≉0 : Core.2≉0 rawPseudocode) where private open module C = RawPseudocode rawPseudocode import Data.Bool as Bool -import Data.Fin as Fin +open import Data.Fin using (fromℕ; suc; inject₁) open import Data.Fin.Patterns -open import Data.Nat using (suc) +open import Data.Nat using (ℕ; suc) open import Data.Vec using (Vec; _∷_) +open import Data.Vec.Relation.Unary.All as All using (All) open import Function using (_∘_) open import Helium.Data.Pseudocode.Core open import Helium.Semantics.Axiomatic.Assertion rawPseudocode as Asrt -open import Helium.Semantics.Axiomatic.Term rawPseudocode as Term using (var; meta; func₀; func₁; 𝒦; ℰ; ℰ′) +open import Helium.Semantics.Axiomatic.Term rawPseudocode as Term using (↓_) open import Level using (_⊔_; lift; lower) renaming (suc to ℓsuc) open import Relation.Nullary.Decidable.Core using (toWitness) -open import Relation.Unary using (_⊆′_) - -module _ (2≉0 : Term.2≉0) {o} {Σ : Vec Type o} where - open Code Σ - data HoareTriple {n} {Γ : Vec Type n} {m} {Δ : Vec Type m} : Assertion Σ Γ Δ → Statement Γ → Assertion Σ Γ Δ → Set (ℓsuc (b₁ ⊔ i₁ ⊔ r₁)) where - _∙_ : ∀ {P Q R s₁ s₂} → HoareTriple P s₁ Q → HoareTriple Q s₂ R → HoareTriple P (s₁ ∙ s₂) R - skip : ∀ {P} → HoareTriple P skip P - - assign : ∀ {P t ref canAssign e} → HoareTriple (subst 2≉0 P (toWitness canAssign) (ℰ 2≉0 e)) (_≔_ {t = t} ref {canAssign} e) P - declare : ∀ {t P Q e s} → HoareTriple (P ∧ equal (var 0F) (Term.wknVar (ℰ 2≉0 e))) s (Asrt.wknVar Q) → HoareTriple (Asrt.elimVar P (ℰ 2≉0 e)) (declare {t = t} e s) Q - invoke : ∀ {m ts P Q s es} → HoareTriple P s (Asrt.addVars Q) → HoareTriple (Asrt.substVars P (ℰ′ 2≉0 es)) (invoke {m = m} {ts} (s ∙end) es) (Asrt.addVars Q) - if : ∀ {P Q e s} → HoareTriple (P ∧ equal (ℰ 2≉0 e) (𝒦 (Bool.true ′b))) s Q → HoareTriple (P ∧ equal (ℰ 2≉0 e) (𝒦 (Bool.false ′b))) skip Q → HoareTriple P (if e then s) Q - if-else : ∀ {P Q e s₁ s₂} → HoareTriple (P ∧ equal (ℰ 2≉0 e) (𝒦 (Bool.true ′b))) s₁ Q → HoareTriple (P ∧ equal (ℰ 2≉0 e) (𝒦 (Bool.false ′b))) s₂ Q → HoareTriple P (if e then s₁ else s₂) Q - for : ∀ {m} {I : Assertion Σ Γ (fin (suc m) ∷ Δ)} {s} → HoareTriple (some (Asrt.wknVar (Asrt.wknMetaAt 1F I) ∧ equal (meta 1F) (var 0F) ∧ equal (meta 0F) (func₁ (lift ∘ Fin.inject₁ ∘ lower) (meta 1F)))) s (some (Asrt.wknVar (Asrt.wknMetaAt 1F I) ∧ equal (meta 0F) (func₁ (lift ∘ Fin.suc ∘ lower) (meta 1F)))) → HoareTriple (some (I ∧ equal (meta 0F) (func₀ (lift 0F)))) (for m s) (some (I ∧ equal (meta 0F) (func₀ (lift (Fin.fromℕ m))))) - - consequence : ∀ {P₁ P₂ Q₁ Q₂ s} → (∀ σ γ δ → ⟦ P₁ ⟧ σ γ δ → ⟦ P₂ ⟧ σ γ δ) → HoareTriple P₂ s Q₂ → (∀ σ γ δ → ⟦ Q₂ ⟧ σ γ δ → ⟦ Q₁ ⟧ σ γ δ) → HoareTriple P₁ s Q₁ - some : ∀ {t P Q s} → HoareTriple P s Q → HoareTriple (some {t = t} P) s (some Q) + +open Term.Term +open Semantics 2≉0 + +private + variable + i j k m n : ℕ + t : Type + Σ Γ Δ ts : Vec Type n + P Q R S : Assertion Σ Γ Δ + ref : Reference Σ Γ t + e val : Expression Σ Γ t + es : All (Expression Σ Γ) ts + s s₁ s₂ : Statement Σ Γ + + ℓ = b₁ ⊔ i₁ ⊔ r₁ + +infix 4 _⊆_ +record _⊆_ (P : Assertion Σ Γ Δ) (Q : Assertion Σ Γ Δ) : Set ℓ where + constructor arr + field + implies : ∀ σ γ δ → ⟦ P ⟧ σ γ δ → ⟦ Q ⟧ σ γ δ + +open _⊆_ public + +data HoareTriple {Σ : Vec Type i} {Γ : Vec Type j} {Δ : Vec Type k} : Assertion Σ Γ Δ → Statement Σ Γ → Assertion Σ Γ Δ → Set (ℓsuc (b₁ ⊔ i₁ ⊔ r₁)) where + seq : ∀ Q → HoareTriple P s Q → HoareTriple Q s₁ R → HoareTriple P (s ∙ s₁) R + skip : P ⊆ Q → HoareTriple P skip Q + assign : subst P ref (↓ val) ⊆ Q → HoareTriple P (ref ≔ val) Q + declare : HoareTriple (Var.weaken 0F P ∧ equal (var 0F) (Term.Var.weaken 0F (↓ e))) s (Var.weaken 0F Q) → HoareTriple P (declare e s) Q + invoke : ∀ (Q R : Assertion Σ ts Δ) → P ⊆ Var.elimAll Q (All.map ↓_ es) → HoareTriple Q s R → Var.inject Γ R ⊆ Var.raise ts S → HoareTriple P (invoke (s ∙end) es) S + if : HoareTriple (P ∧ pred (↓ e)) s Q → P ∧ pred (↓ inv e) ⊆ Q → HoareTriple P (if e then s) Q + if-else : HoareTriple (P ∧ pred (↓ e)) s Q → HoareTriple (P ∧ pred (↓ inv e)) s Q → HoareTriple P (if e then s) Q + for : ∀ (I : Assertion _ _ (fin _ ∷ _)) → P ⊆ Meta.elim 0F I (↓ lit 0F) → HoareTriple {Δ = fin _ ∷ Δ} (Var.weaken 0F (Meta.elim 1F (Meta.weaken 0F I) (fin inject₁ (cons (meta 0F) nil)))) s (Var.weaken 0F (Meta.elim 1F (Meta.weaken 0F I) (fin suc (cons (meta 0F) nil)))) → Meta.elim 0F I (↓ lit (fromℕ m)) ⊆ Q → HoareTriple P (for m s) Q + some : HoareTriple P s Q → HoareTriple (some P) s (some Q) diff --git a/src/Helium/Semantics/Core.agda b/src/Helium/Semantics/Core.agda new file mode 100644 index 0000000..688f6f6 --- /dev/null +++ b/src/Helium/Semantics/Core.agda @@ -0,0 +1,209 @@ +------------------------------------------------------------------------ +-- Agda Helium +-- +-- Base definitions for semantics +------------------------------------------------------------------------ + +{-# OPTIONS --safe --without-K #-} + +open import Helium.Data.Pseudocode.Algebra using (RawPseudocode) + +module Helium.Semantics.Core + {b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃} + (rawPseudocode : RawPseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) + where + +private + open module C = RawPseudocode rawPseudocode + +open import Algebra.Core using (Op₁) +open import Data.Bool as Bool using (Bool) +open import Data.Fin as Fin using (Fin; zero; suc) +open import Data.Fin.Patterns +import Data.Fin.Properties as Finₚ +open import Data.Integer as 𝕀 using () renaming (ℤ to 𝕀) +open import Data.Nat as ℕ using (ℕ; suc) +import Data.Nat.Properties as ℕₚ +open import Data.Product as × using (_×_; _,_) +open import Data.Sign using (Sign) +open import Data.Unit using (⊤) +open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup; map; take; drop) +open import Data.Vec.Relation.Binary.Pointwise.Extensional using (Pointwise; decidable) +open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) +open import Function +open import Helium.Data.Pseudocode.Core +open import Level hiding (suc) +open import Relation.Binary +import Relation.Binary.Construct.On as On +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary using (¬_) +open import Relation.Nullary.Decidable.Core using (map′) + +private + variable + a : Level + A : Set a + t t′ t₁ t₂ : Type + m n : ℕ + Γ Δ Σ ts : Vec Type m + + ℓ = b₁ ⊔ i₁ ⊔ r₁ + ℓ₁ = ℓ ⊔ b₂ ⊔ i₂ ⊔ r₂ + ℓ₂ = i₁ ⊔ i₃ ⊔ r₁ ⊔ r₃ + + Sign⇒- : Sign → Op₁ A → Op₁ A + Sign⇒- Sign.+ f = id + Sign⇒- Sign.- f = f + +open ℕₚ.≤-Reasoning + +𝕀⇒ℤ : 𝕀 → ℤ +𝕀⇒ℤ z = Sign⇒- (𝕀.sign z) ℤ.-_ (𝕀.∣ z ∣ ℤ′.×′ 1ℤ) + +𝕀⇒ℝ : 𝕀 → ℝ +𝕀⇒ℝ z = Sign⇒- (𝕀.sign z) ℝ.-_ (𝕀.∣ z ∣ ℝ′.×′ 1ℝ) + +castVec : .(eq : m ≡ n) → Vec A m → Vec A n +castVec {m = .0} {0} eq [] = [] +castVec {m = .suc m} {suc n} eq (x ∷ xs) = x ∷ castVec (ℕₚ.suc-injective eq) xs + +⟦_⟧ₜ : Type → Set ℓ +⟦_⟧ₜ′ : Vec Type n → Set ℓ + +⟦ bool ⟧ₜ = Lift ℓ Bool +⟦ int ⟧ₜ = Lift ℓ ℤ +⟦ fin n ⟧ₜ = Lift ℓ (Fin n) +⟦ real ⟧ₜ = Lift ℓ ℝ +⟦ bit ⟧ₜ = Lift ℓ Bit +⟦ tuple ts ⟧ₜ = ⟦ ts ⟧ₜ′ +⟦ array t n ⟧ₜ = Vec ⟦ t ⟧ₜ n + +⟦ [] ⟧ₜ′ = Lift ℓ ⊤ +⟦ t ∷ [] ⟧ₜ′ = ⟦ t ⟧ₜ +⟦ t ∷ t₁ ∷ ts ⟧ₜ′ = ⟦ t ⟧ₜ × ⟦ t₁ ∷ ts ⟧ₜ′ + +fetch : ∀ (i : Fin n) Γ → ⟦ Γ ⟧ₜ′ → ⟦ lookup Γ i ⟧ₜ +fetch 0F (t ∷ []) x = x +fetch 0F (t ∷ t₁ ∷ Γ) (x , xs) = x +fetch (suc i) (t ∷ t₁ ∷ Γ) (x , xs) = fetch i (t₁ ∷ Γ) xs + +updateAt : ∀ (i : Fin n) Γ → ⟦ lookup Γ i ⟧ₜ → ⟦ Γ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ +updateAt 0F (t ∷ []) v x = v +updateAt 0F (t ∷ t₁ ∷ Γ) v (x , xs) = v , xs +updateAt (suc i) (t ∷ t₁ ∷ Γ) v (x , xs) = x , updateAt i (t₁ ∷ Γ) v xs + +cons′ : ∀ (ts : Vec Type n) → ⟦ t ⟧ₜ → ⟦ tuple ts ⟧ₜ → ⟦ tuple (t ∷ ts) ⟧ₜ +cons′ [] x xs = x +cons′ (_ ∷ _) x xs = x , xs + +head′ : ∀ (ts : Vec Type n) → ⟦ tuple (t ∷ ts) ⟧ₜ → ⟦ t ⟧ₜ +head′ [] x = x +head′ (_ ∷ _) (x , xs) = x + +tail′ : ∀ (ts : Vec Type n) → ⟦ tuple (t ∷ ts) ⟧ₜ → ⟦ tuple ts ⟧ₜ +tail′ [] x = _ +tail′ (_ ∷ _) (x , xs) = xs + +_≈_ : ⦃ HasEquality t ⦄ → Rel ⟦ t ⟧ₜ ℓ₁ +_≈_ ⦃ bool ⦄ = Lift ℓ₁ ∘₂ _≡_ on lower +_≈_ ⦃ int ⦄ = Lift ℓ₁ ∘₂ ℤ._≈_ on lower +_≈_ ⦃ fin ⦄ = Lift ℓ₁ ∘₂ _≡_ on lower +_≈_ ⦃ real ⦄ = Lift ℓ₁ ∘₂ ℝ._≈_ on lower +_≈_ ⦃ bit ⦄ = Lift ℓ₁ ∘₂ 𝔹._≈_ on lower +_≈_ ⦃ array ⦄ = Pointwise _≈_ + +_<_ : ⦃ Ordered t ⦄ → Rel ⟦ t ⟧ₜ ℓ₂ +_<_ ⦃ int ⦄ = Lift ℓ₂ ∘₂ ℤ._<_ on lower +_<_ ⦃ fin ⦄ = Lift ℓ₂ ∘₂ Fin._<_ on lower +_<_ ⦃ real ⦄ = Lift ℓ₂ ∘₂ ℝ._<_ on lower + +≈-dec : ⦃ hasEq : HasEquality t ⦄ → Decidable (_≈_ ⦃ hasEq ⦄) +≈-dec ⦃ bool ⦄ = map′ lift lower ∘₂ On.decidable lower _≡_ Bool._≟_ +≈-dec ⦃ int ⦄ = map′ lift lower ∘₂ On.decidable lower ℤ._≈_ _≟ᶻ_ +≈-dec ⦃ fin ⦄ = map′ lift lower ∘₂ On.decidable lower _≡_ Fin._≟_ +≈-dec ⦃ real ⦄ = map′ lift lower ∘₂ On.decidable lower ℝ._≈_ _≟ʳ_ +≈-dec ⦃ bit ⦄ = map′ lift lower ∘₂ On.decidable lower 𝔹._≈_ _≟ᵇ₁_ +≈-dec ⦃ array ⦄ = decidable ≈-dec + +<-dec : ⦃ ordered : Ordered t ⦄ → Decidable (_<_ ⦃ ordered ⦄) +<-dec ⦃ int ⦄ = map′ lift lower ∘₂ On.decidable lower ℤ._<_ _<ᶻ?_ +<-dec ⦃ fin ⦄ = map′ lift lower ∘₂ On.decidable lower Fin._<_ Fin._<?_ +<-dec ⦃ real ⦄ = map′ lift lower ∘₂ On.decidable lower ℝ._<_ _<ʳ?_ + +Κ[_]_ : ∀ t → literalType t → ⟦ t ⟧ₜ +Κ[ bool ] x = lift x +Κ[ int ] x = lift (𝕀⇒ℤ x) +Κ[ fin n ] x = lift x +Κ[ real ] x = lift (𝕀⇒ℝ x) +Κ[ bit ] x = lift (Bool.if x then 1𝔹 else 0𝔹) +Κ[ tuple [] ] x = _ +Κ[ tuple (t ∷ []) ] x = Κ[ t ] x +Κ[ tuple (t ∷ t₁ ∷ ts) ] (x , xs) = Κ[ t ] x , Κ[ tuple (t₁ ∷ ts) ] xs +Κ[ array t n ] x = map Κ[ t ]_ x + +2≉0 : Set _ +2≉0 = ¬ 2 ℝ′.×′ 1ℝ ℝ.≈ 0ℝ + +neg : ⦃ IsNumeric t ⦄ → Op₁ ⟦ t ⟧ₜ +neg ⦃ int ⦄ = lift ∘ ℤ.-_ ∘ lower +neg ⦃ real ⦄ = lift ∘ ℝ.-_ ∘ lower + +add : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → ⟦ t₁ ⟧ₜ → ⟦ t₂ ⟧ₜ → ⟦ isNum₁ +ᵗ isNum₂ ⟧ₜ +add ⦃ int ⦄ ⦃ int ⦄ x y = lift (lower x ℤ.+ lower y) +add ⦃ int ⦄ ⦃ real ⦄ x y = lift (lower x /1 ℝ.+ lower y) +add ⦃ real ⦄ ⦃ int ⦄ x y = lift (lower x ℝ.+ lower y /1) +add ⦃ real ⦄ ⦃ real ⦄ x y = lift (lower x ℝ.+ lower y) + +mul : ⦃ isNum₁ : IsNumeric t₁ ⦄ → ⦃ isNum₂ : IsNumeric t₂ ⦄ → ⟦ t₁ ⟧ₜ → ⟦ t₂ ⟧ₜ → ⟦ isNum₁ +ᵗ isNum₂ ⟧ₜ +mul ⦃ int ⦄ ⦃ int ⦄ x y = lift (lower x ℤ.* lower y) +mul ⦃ int ⦄ ⦃ real ⦄ x y = lift (lower x /1 ℝ.* lower y) +mul ⦃ real ⦄ ⦃ int ⦄ x y = lift (lower x ℝ.* lower y /1) +mul ⦃ real ⦄ ⦃ real ⦄ x y = lift (lower x ℝ.* lower y) + +pow : ⦃ IsNumeric t ⦄ → ⟦ t ⟧ₜ → ℕ → ⟦ t ⟧ₜ +pow ⦃ int ⦄ = lift ∘₂ ℤ′._^′_ ∘ lower +pow ⦃ real ⦄ = lift ∘₂ ℝ′._^′_ ∘ lower + +shift : 2≉0 → ℤ → ℕ → ℤ +shift 2≉0 z n = ⌊ z /1 ℝ.* 2≉0 ℝ.⁻¹ ℝ′.^′ n ⌋ + +lowerFin : ∀ (ms : Vec ℕ n) → ⟦ tuple (map fin ms) ⟧ₜ → literalTypes (map fin ms) +lowerFin [] _ = _ +lowerFin (_ ∷ []) x = lower x +lowerFin (_ ∷ m₁ ∷ ms) (x , xs) = lower x , lowerFin (m₁ ∷ ms) xs + +mergeVec : Vec A m → Vec A n → Fin (suc n) → Vec A (n ℕ.+ m) +mergeVec {m = m} {n} xs ys i = castVec eq (low ++ xs ++ high) + where + i′ = Fin.toℕ i + ys′ = castVec (sym (ℕₚ.m+[n∸m]≡n (ℕ.≤-pred (Finₚ.toℕ<n i)))) ys + low = take i′ ys′ + high = drop i′ ys′ + eq = begin-equality + i′ ℕ.+ (m ℕ.+ (n ℕ.∸ i′)) ≡⟨ ℕₚ.+-comm i′ _ ⟩ + m ℕ.+ (n ℕ.∸ i′) ℕ.+ i′ ≡⟨ ℕₚ.+-assoc m _ _ ⟩ + m ℕ.+ (n ℕ.∸ i′ ℕ.+ i′) ≡⟨ cong (m ℕ.+_) (ℕₚ.m∸n+n≡m (ℕ.≤-pred (Finₚ.toℕ<n i))) ⟩ + m ℕ.+ n ≡⟨ ℕₚ.+-comm m n ⟩ + n ℕ.+ m ∎ + +sliceVec : Vec A (n ℕ.+ m) → Fin (suc n) → Vec A m +sliceVec {n = n} {m} xs i = take m (drop i′ (castVec eq xs)) + where + i′ = Fin.toℕ i + eq = sym $ begin-equality + i′ ℕ.+ (m ℕ.+ (n ℕ.∸ i′)) ≡⟨ ℕₚ.+-comm i′ _ ⟩ + m ℕ.+ (n ℕ.∸ i′) ℕ.+ i′ ≡⟨ ℕₚ.+-assoc m _ _ ⟩ + m ℕ.+ (n ℕ.∸ i′ ℕ.+ i′) ≡⟨ cong (m ℕ.+_) (ℕₚ.m∸n+n≡m (ℕ.≤-pred (Finₚ.toℕ<n i))) ⟩ + m ℕ.+ n ≡⟨ ℕₚ.+-comm m n ⟩ + n ℕ.+ m ∎ + +cutVec : Vec A (n ℕ.+ m) → Fin (suc n) → Vec A n +cutVec {n = n} {m} xs i = castVec (ℕₚ.m+[n∸m]≡n (ℕ.≤-pred (Finₚ.toℕ<n i))) (take i′ (castVec eq xs) ++ drop m (drop i′ (castVec eq xs))) + where + i′ = Fin.toℕ i + eq = sym $ begin-equality + i′ ℕ.+ (m ℕ.+ (n ℕ.∸ i′)) ≡⟨ ℕₚ.+-comm i′ _ ⟩ + m ℕ.+ (n ℕ.∸ i′) ℕ.+ i′ ≡⟨ ℕₚ.+-assoc m _ _ ⟩ + m ℕ.+ (n ℕ.∸ i′ ℕ.+ i′) ≡⟨ cong (m ℕ.+_) (ℕₚ.m∸n+n≡m (ℕ.≤-pred (Finₚ.toℕ<n i))) ⟩ + m ℕ.+ n ≡⟨ ℕₚ.+-comm m n ⟩ + n ℕ.+ m ∎ diff --git a/src/Helium/Semantics/Denotational/Core.agda b/src/Helium/Semantics/Denotational/Core.agda index 07c71bd..c015cbc 100644 --- a/src/Helium/Semantics/Denotational/Core.agda +++ b/src/Helium/Semantics/Denotational/Core.agda @@ -16,284 +16,144 @@ module Helium.Semantics.Denotational.Core private open module C = RawPseudocode rawPseudocode -open import Data.Bool as Bool using (Bool; true; false) -open import Data.Fin as Fin using (Fin; zero; suc) -import Data.Fin.Properties as Finₚ -open import Data.Nat as ℕ using (ℕ; zero; suc) -import Data.Nat.Properties as ℕₚ -open import Data.Product as P using (_×_; _,_) -open import Data.Sum as S using (_⊎_) renaming (inj₁ to next; inj₂ to ret) -open import Data.Unit using (⊤) -open import Data.Vec as Vec using (Vec; []; _∷_) +import Data.Bool as Bool +open import Data.Empty using (⊥-elim) +import Data.Fin as Fin +import Data.Integer as 𝕀 +open import Data.Nat using (ℕ) +open import Data.Product using (_×_; _,_; proj₁; proj₂; <_,_>; uncurry) +open import Data.Vec as Vec using (Vec; []; _∷_; map; zipWith) open import Data.Vec.Relation.Unary.All using (All; []; _∷_) -import Data.Vec.Functional as VecF -open import Function using (case_of_; _∘′_; id) +open import Function open import Helium.Data.Pseudocode.Core -import Induction as I -import Induction.WellFounded as Wf -open import Level using (Level; _⊔_; 0ℓ) -open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; module ≡-Reasoning) -open import Relation.Nullary using (does) renaming (¬_ to ¬′_) -open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromWitness) +open import Helium.Semantics.Core rawPseudocode +open import Level +open import Relation.Binary.PropositionalEquality using (sym) +open import Relation.Nullary using (does) -⟦_⟧ₗ : Type → Level -⟦ bool ⟧ₗ = 0ℓ -⟦ int ⟧ₗ = i₁ -⟦ fin n ⟧ₗ = 0ℓ -⟦ real ⟧ₗ = r₁ -⟦ bit ⟧ₗ = b₁ -⟦ bits n ⟧ₗ = b₁ -⟦ tuple n ts ⟧ₗ = helper ts - where - helper : ∀ {n} → Vec Type n → Level - helper [] = 0ℓ - helper (t ∷ ts) = ⟦ t ⟧ₗ ⊔ helper ts -⟦ array t n ⟧ₗ = ⟦ t ⟧ₗ - -⟦_⟧ₜ : ∀ t → Set ⟦ t ⟧ₗ -⟦_⟧ₜ′ : ∀ {n} ts → Set ⟦ tuple n ts ⟧ₗ - -⟦ bool ⟧ₜ = Bool -⟦ int ⟧ₜ = ℤ -⟦ fin n ⟧ₜ = Fin n -⟦ real ⟧ₜ = ℝ -⟦ bit ⟧ₜ = Bit -⟦ bits n ⟧ₜ = Bits n -⟦ tuple n ts ⟧ₜ = ⟦ ts ⟧ₜ′ -⟦ array t n ⟧ₜ = Vec ⟦ t ⟧ₜ n - -⟦ [] ⟧ₜ′ = ⊤ -⟦ t ∷ [] ⟧ₜ′ = ⟦ t ⟧ₜ -⟦ t ∷ t′ ∷ ts ⟧ₜ′ = ⟦ t ⟧ₜ × ⟦ t′ ∷ ts ⟧ₜ′ - --- The case for bitvectors looks odd so that the right-most bit is bit 0. -𝒦 : ∀ {t} → Literal t → ⟦ t ⟧ₜ -𝒦 (x ′b) = x -𝒦 (x ′i) = x ℤ′.×′ 1ℤ -𝒦 (x ′f) = x -𝒦 (x ′r) = x ℝ′.×′ 1ℝ -𝒦 (x ′x) = Bool.if x then 1𝔹 else 0𝔹 -𝒦 (xs ′xs) = Vec.foldl Bits (λ bs b → (Bool.if b then 1𝔹 else 0𝔹) VecF.∷ bs) VecF.[] xs -𝒦 (x ′a) = Vec.replicate (𝒦 x) - -fetch : ∀ {n} ts → ⟦ tuple n ts ⟧ₜ → ∀ i → ⟦ Vec.lookup ts i ⟧ₜ -fetch (_ ∷ []) x zero = x -fetch (_ ∷ _ ∷ _) (x , xs) zero = x -fetch (_ ∷ t ∷ ts) (x , xs) (suc i) = fetch (t ∷ ts) xs i - -updateAt : ∀ {n} ts i → ⟦ Vec.lookup ts i ⟧ₜ → ⟦ tuple n ts ⟧ₜ → ⟦ tuple n ts ⟧ₜ -updateAt (_ ∷ []) zero v _ = v -updateAt (_ ∷ _ ∷ _) zero v (_ , xs) = v , xs -updateAt (_ ∷ t ∷ ts) (suc i) v (x , xs) = x , updateAt (t ∷ ts) i v xs - -tupCons : ∀ {n t} ts → ⟦ t ⟧ₜ → ⟦ tuple n ts ⟧ₜ → ⟦ tuple _ (t ∷ ts) ⟧ₜ -tupCons [] x xs = x -tupCons (t ∷ ts) x xs = x , xs - -tupHead : ∀ {n t} ts → ⟦ tuple (suc n) (t ∷ ts) ⟧ₜ → ⟦ t ⟧ₜ -tupHead {t = t} ts xs = fetch (t ∷ ts) xs zero - -tupTail : ∀ {n t} ts → ⟦ tuple _ (t ∷ ts) ⟧ₜ → ⟦ tuple n ts ⟧ₜ -tupTail [] x = _ -tupTail (_ ∷ _) (x , xs) = xs - -equal : ∀ {t} → HasEquality t → ⟦ t ⟧ₜ → ⟦ t ⟧ₜ → Bool -equal bool x y = does (x Bool.≟ y) -equal int x y = does (x ≟ᶻ y) -equal fin x y = does (x Fin.≟ y) -equal real x y = does (x ≟ʳ y) -equal bit x y = does (x ≟ᵇ₁ y) -equal bits x y = does (x ≟ᵇ y) - -comp : ∀ {t} → IsNumeric t → ⟦ t ⟧ₜ → ⟦ t ⟧ₜ → Bool -comp int x y = does (x <ᶻ? y) -comp real x y = does (x <ʳ? y) - --- 0 of y is 0 of result -join : ∀ t {m n} → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ → ⟦ asType t (n ℕ.+ m) ⟧ₜ -join bits x y = y VecF.++ x -join (array _) x y = y Vec.++ x - --- take from 0 -take : ∀ t i {j} → ⟦ asType t (i ℕ.+ j) ⟧ₜ → ⟦ asType t i ⟧ₜ -take bits i x = VecF.take i x -take (array _) i x = Vec.take i x - --- drop from 0 -drop : ∀ t i {j} → ⟦ asType t (i ℕ.+ j) ⟧ₜ → ⟦ asType t j ⟧ₜ -drop bits i x = VecF.drop i x -drop (array _) i x = Vec.drop i x - -casted : ∀ t {i j} → .(eq : i ≡ j) → ⟦ asType t i ⟧ₜ → ⟦ asType t j ⟧ₜ -casted bits eq x = x ∘′ Fin.cast (≡.sym eq) -casted (array _) {j = zero} eq [] = [] -casted (array t) {j = suc _} eq (x ∷ y) = x ∷ casted (array t) (ℕₚ.suc-injective eq) y - -module _ where - m≤n⇒m+k≡n : ∀ {m n} → m ℕ.≤ n → P.∃ λ k → m ℕ.+ k ≡ n - m≤n⇒m+k≡n ℕ.z≤n = _ , ≡.refl - m≤n⇒m+k≡n (ℕ.s≤s m≤n) = P.dmap id (≡.cong suc) (m≤n⇒m+k≡n m≤n) - - slicedSize : ∀ n m (i : Fin (suc n)) → P.∃ λ k → n ℕ.+ m ≡ Fin.toℕ i ℕ.+ (m ℕ.+ k) × Fin.toℕ i ℕ.+ k ≡ n - slicedSize n m i = k , (begin - n ℕ.+ m ≡˘⟨ ≡.cong (ℕ._+ m) (P.proj₂ i+k≡n) ⟩ - (Fin.toℕ i ℕ.+ k) ℕ.+ m ≡⟨ ℕₚ.+-assoc (Fin.toℕ i) k m ⟩ - Fin.toℕ i ℕ.+ (k ℕ.+ m) ≡⟨ ≡.cong (Fin.toℕ i ℕ.+_) (ℕₚ.+-comm k m) ⟩ - Fin.toℕ i ℕ.+ (m ℕ.+ k) ∎) , - P.proj₂ i+k≡n - where - open ≡-Reasoning - i+k≡n = m≤n⇒m+k≡n (ℕₚ.≤-pred (Finₚ.toℕ<n i)) - k = P.proj₁ i+k≡n - - -- 0 of x is i of result - spliced : ∀ t {m n} → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ → ⟦ fin (suc n) ⟧ₜ → ⟦ asType t (n ℕ.+ m) ⟧ₜ - spliced t {m} x y i = casted t eq (join t (join t high x) low) - where - reasoning = slicedSize _ m i - k = P.proj₁ reasoning - n≡i+k = ≡.sym (P.proj₂ (P.proj₂ reasoning)) - low = take t (Fin.toℕ i) (casted t n≡i+k y) - high = drop t (Fin.toℕ i) (casted t n≡i+k y) - eq = ≡.sym (P.proj₁ (P.proj₂ reasoning)) - - sliced : ∀ t {m n} → ⟦ asType t (n ℕ.+ m) ⟧ₜ → ⟦ fin (suc n) ⟧ₜ → ⟦ asType t m ∷ asType t n ∷ [] ⟧ₜ′ - sliced t {m} x i = middle , casted t i+k≡n (join t high low) - where - reasoning = slicedSize _ m i - k = P.proj₁ reasoning - i+k≡n = P.proj₂ (P.proj₂ reasoning) - eq = P.proj₁ (P.proj₂ reasoning) - low = take t (Fin.toℕ i) (casted t eq x) - middle = take t m (drop t (Fin.toℕ i) (casted t eq x)) - high = drop t m (drop t (Fin.toℕ i) (casted t eq x)) - -box : ∀ t → ⟦ elemType t ⟧ₜ → ⟦ asType t 1 ⟧ₜ -box bits v = v VecF.∷ VecF.[] -box (array t) v = v ∷ [] - -unboxed : ∀ t → ⟦ asType t 1 ⟧ₜ → ⟦ elemType t ⟧ₜ -unboxed bits v = v (Fin.zero) -unboxed (array t) (v ∷ []) = v - -neg : ∀ {t} → IsNumeric t → ⟦ t ⟧ₜ → ⟦ t ⟧ₜ -neg int x = ℤ.- x -neg real x = ℝ.- x - -add : ∀ {t₁ t₂} (isNum₁ : True (isNumeric? t₁)) (isNum₂ : True (isNumeric? t₂)) → ⟦ t₁ ⟧ₜ → ⟦ t₂ ⟧ₜ → ⟦ combineNumeric t₁ t₂ isNum₁ isNum₂ ⟧ₜ -add {t₁ = int} {t₂ = int} _ _ x y = x ℤ.+ y -add {t₁ = int} {t₂ = real} _ _ x y = x /1 ℝ.+ y -add {t₁ = real} {t₂ = int} _ _ x y = x ℝ.+ y /1 -add {t₁ = real} {t₂ = real} _ _ x y = x ℝ.+ y - -mul : ∀ {t₁ t₂} (isNum₁ : True (isNumeric? t₁)) (isNum₂ : True (isNumeric? t₂)) → ⟦ t₁ ⟧ₜ → ⟦ t₂ ⟧ₜ → ⟦ combineNumeric t₁ t₂ isNum₁ isNum₂ ⟧ₜ -mul {t₁ = int} {t₂ = int} _ _ x y = x ℤ.* y -mul {t₁ = int} {t₂ = real} _ _ x y = x /1 ℝ.* y -mul {t₁ = real} {t₂ = int} _ _ x y = x ℝ.* y /1 -mul {t₁ = real} {t₂ = real} _ _ x y = x ℝ.* y - -pow : ∀ {t} → IsNumeric t → ⟦ t ⟧ₜ → ℕ → ⟦ t ⟧ₜ -pow int x n = x ℤ′.^′ n -pow real x n = x ℝ′.^′ n - -shiftr : ¬′ 2 ℝ′.×′ 1ℝ ℝ.≈ 0ℝ → ⟦ int ⟧ₜ → ℕ → ⟦ int ⟧ₜ -shiftr 2≉0 x n = ⌊ x /1 ℝ.* 2≉0 ℝ.⁻¹ ℝ′.^′ n ⌋ - -module Expression - {o} {Σ : Vec Type o} - (2≉0 : ¬′ 2 ℝ′.×′ 1ℝ ℝ.≈ 0ℝ) - where - - open Code Σ - - ⟦_⟧ᵉ : ∀ {n} {Γ : Vec Type n} {t} → Expression Γ t → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ t ⟧ₜ - ⟦_⟧ˢ : ∀ {n} {Γ : Vec Type n} → Statement Γ → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ - ⟦_⟧ᶠ : ∀ {n} {Γ : Vec Type n} {ret} → Function Γ ret → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ ret ⟧ₜ - ⟦_⟧ᵖ : ∀ {n} {Γ : Vec Type n} → Procedure Γ → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ - ⟦_⟧ᵉ′ : ∀ {n} {Γ : Vec Type n} {m ts} → All (Expression Γ) ts → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ tuple m ts ⟧ₜ - update : ∀ {n Γ t e} → CanAssign {n} {Γ} {t} e → ⟦ t ⟧ₜ → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ - - ⟦ lit x ⟧ᵉ σ γ = 𝒦 x - ⟦ state i ⟧ᵉ σ γ = fetch Σ σ i - ⟦_⟧ᵉ {Γ = Γ} (var i) σ γ = fetch Γ γ i - ⟦ abort e ⟧ᵉ σ γ = case ⟦ e ⟧ᵉ σ γ of λ () - ⟦ _≟_ {hasEquality = hasEq} e e₁ ⟧ᵉ σ γ = equal (toWitness hasEq) (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - ⟦ _<?_ {isNumeric = isNum} e e₁ ⟧ᵉ σ γ = comp (toWitness isNum) (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - ⟦ inv e ⟧ᵉ σ γ = Bool.not (⟦ e ⟧ᵉ σ γ) - ⟦ e && e₁ ⟧ᵉ σ γ = Bool.if ⟦ e ⟧ᵉ σ γ then ⟦ e₁ ⟧ᵉ σ γ else false - ⟦ e || e₁ ⟧ᵉ σ γ = Bool.if ⟦ e ⟧ᵉ σ γ then true else ⟦ e₁ ⟧ᵉ σ γ - ⟦ not e ⟧ᵉ σ γ = Bits.¬_ (⟦ e ⟧ᵉ σ γ) - ⟦ e and e₁ ⟧ᵉ σ γ = ⟦ e ⟧ᵉ σ γ Bits.∧ ⟦ e₁ ⟧ᵉ σ γ - ⟦ e or e₁ ⟧ᵉ σ γ = ⟦ e ⟧ᵉ σ γ Bits.∨ ⟦ e₁ ⟧ᵉ σ γ - ⟦ [_] {t = t} e ⟧ᵉ σ γ = box t (⟦ e ⟧ᵉ σ γ) - ⟦ unbox {t = t} e ⟧ᵉ σ γ = unboxed t (⟦ e ⟧ᵉ σ γ) - ⟦ splice {t = t} e e₁ e₂ ⟧ᵉ σ γ = spliced t (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) (⟦ e₂ ⟧ᵉ σ γ) - ⟦ cut {t = t} e e₁ ⟧ᵉ σ γ = sliced t (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - ⟦ cast {t = t} eq e ⟧ᵉ σ γ = casted t eq (⟦ e ⟧ᵉ σ γ) - ⟦ -_ {isNumeric = isNum} e ⟧ᵉ σ γ = neg (toWitness isNum) (⟦ e ⟧ᵉ σ γ) - ⟦ _+_ {isNumeric₁ = isNum₁} {isNumeric₂ = isNum₂} e e₁ ⟧ᵉ σ γ = add isNum₁ isNum₂ (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - ⟦ _*_ {isNumeric₁ = isNum₁} {isNumeric₂ = isNum₂} e e₁ ⟧ᵉ σ γ = mul isNum₁ isNum₂ (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - -- ⟦ e / e₁ ⟧ᵉ σ γ = {!!} - ⟦ _^_ {isNumeric = isNum} e n ⟧ᵉ σ γ = pow (toWitness isNum) (⟦ e ⟧ᵉ σ γ) n - ⟦ _>>_ e n ⟧ᵉ σ γ = shiftr 2≉0 (⟦ e ⟧ᵉ σ γ) n - ⟦ rnd e ⟧ᵉ σ γ = ⌊ ⟦ e ⟧ᵉ σ γ ⌋ - ⟦ fin x e ⟧ᵉ σ γ = apply x (⟦ e ⟧ᵉ σ γ) - where - apply : ∀ {k ms n} → (All Fin ms → Fin n) → ⟦ Vec.map {n = k} fin ms ⟧ₜ′ → ⟦ fin n ⟧ₜ - apply {zero} {[]} f xs = f [] - apply {suc k} {_ ∷ ms} f xs = - apply (λ x → f (tupHead (Vec.map fin ms) xs ∷ x)) (tupTail (Vec.map fin ms) xs) - ⟦ asInt e ⟧ᵉ σ γ = Fin.toℕ (⟦ e ⟧ᵉ σ γ) ℤ′.×′ 1ℤ - ⟦ nil ⟧ᵉ σ γ = _ - ⟦ cons {ts = ts} e e₁ ⟧ᵉ σ γ = tupCons ts (⟦ e ⟧ᵉ σ γ) (⟦ e₁ ⟧ᵉ σ γ) - ⟦ head {ts = ts} e ⟧ᵉ σ γ = tupHead ts (⟦ e ⟧ᵉ σ γ) - ⟦ tail {ts = ts} e ⟧ᵉ σ γ = tupTail ts (⟦ e ⟧ᵉ σ γ) - ⟦ call f e ⟧ᵉ σ γ = ⟦ f ⟧ᶠ σ (⟦ e ⟧ᵉ′ σ γ) - ⟦ if e then e₁ else e₂ ⟧ᵉ σ γ = Bool.if ⟦ e ⟧ᵉ σ γ then ⟦ e₁ ⟧ᵉ σ γ else ⟦ e₂ ⟧ᵉ σ γ - - ⟦ [] ⟧ᵉ′ σ γ = _ - ⟦ e ∷ [] ⟧ᵉ′ σ γ = ⟦ e ⟧ᵉ σ γ - ⟦ e ∷ e′ ∷ es ⟧ᵉ′ σ γ = ⟦ e ⟧ᵉ σ γ , ⟦ e′ ∷ es ⟧ᵉ′ σ γ - - ⟦ s ∙ s₁ ⟧ˢ σ γ = P.uncurry ⟦ s ⟧ˢ (⟦ s ⟧ˢ σ γ) - ⟦ skip ⟧ˢ σ γ = σ , γ - ⟦ _≔_ ref {canAssign = canAssign} e ⟧ˢ σ γ = update (toWitness canAssign) (⟦ e ⟧ᵉ σ γ) σ γ - ⟦_⟧ˢ {Γ = Γ} (declare e s) σ γ = P.map₂ (tupTail Γ) (⟦ s ⟧ˢ σ (tupCons Γ (⟦ e ⟧ᵉ σ γ) γ)) - ⟦ invoke p e ⟧ˢ σ γ = ⟦ p ⟧ᵖ σ (⟦ e ⟧ᵉ′ σ γ) , γ - ⟦ if e then s₁ ⟧ˢ σ γ = Bool.if ⟦ e ⟧ᵉ σ γ then ⟦ s₁ ⟧ˢ σ γ else (σ , γ) - ⟦ if e then s₁ else s₂ ⟧ˢ σ γ = Bool.if ⟦ e ⟧ᵉ σ γ then ⟦ s₁ ⟧ˢ σ γ else ⟦ s₂ ⟧ˢ σ γ - ⟦_⟧ˢ {Γ = Γ} (for m s) σ γ = helper m ⟦ s ⟧ˢ σ γ - where - helper : ∀ m → (⟦ Σ ⟧ₜ′ → ⟦ fin m ∷ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ fin m ∷ Γ ⟧ₜ′) → ⟦ Σ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ - helper zero s σ γ = σ , γ - helper (suc m) s σ γ = P.uncurry (helper m s′) (P.map₂ (tupTail Γ) (s σ (tupCons Γ zero γ))) - where - s′ : ⟦ Σ ⟧ₜ′ → ⟦ fin m ∷ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ fin m ∷ Γ ⟧ₜ′ - s′ σ γ = - P.map₂ (tupCons Γ (tupHead Γ γ) ∘′ (tupTail Γ)) - (s σ (tupCons Γ (suc (tupHead Γ γ)) (tupTail Γ γ))) - - ⟦ s ∙return e ⟧ᶠ σ γ = P.uncurry ⟦ e ⟧ᵉ (⟦ s ⟧ˢ σ γ) - ⟦_⟧ᶠ {Γ = Γ} (declare e f) σ γ = ⟦ f ⟧ᶠ σ (tupCons Γ (⟦ e ⟧ᵉ σ γ) γ) - - ⟦ s ∙end ⟧ᵖ σ γ = P.proj₁ (⟦ s ⟧ˢ σ γ) - - update (state i) v σ γ = updateAt Σ i v σ , γ - update {Γ = Γ} (var i) v σ γ = σ , updateAt Γ i v γ - update (abort _) v σ γ = σ , γ - update ([_] {t = t} e) v σ γ = update e (unboxed t v) σ γ - update (unbox {t = t} e) v σ γ = update e (box t v) σ γ - update (splice {m = m} {t = t} e e₁ e₂) v σ γ = do - let i = ⟦ e₂ ⟧ᵉ σ γ - let σ′ , γ′ = update e (P.proj₁ (sliced t v i)) σ γ - update e₁ (P.proj₂ (sliced t v i)) σ′ γ′ - update (cut {t = t} a e₂) v σ γ = do - let i = ⟦ e₂ ⟧ᵉ σ γ - update a (spliced t (P.proj₁ v) (P.proj₂ v) i) σ γ - update (cast {t = t} eq e) v σ γ = update e (casted t (≡.sym eq) v) σ γ - update nil v σ γ = σ , γ - update (cons {ts = ts} e e₁) vs σ γ = do - let σ′ , γ′ = update e (tupHead ts vs) σ γ - update e₁ (tupTail ts vs) σ′ γ′ - update (head {ts = ts} {e = e} a) v σ γ = update a (tupCons ts v (tupTail ts (⟦ e ⟧ᵉ σ γ))) σ γ - update (tail {ts = ts} {e = e} a) v σ γ = update a (tupCons ts (tupHead ts (⟦ e ⟧ᵉ σ γ)) v) σ γ +private + variable + n : ℕ + t : Type + Σ Γ ts : Vec Type n + + +module Semantics (2≉0 : 2≉0) where + expr : Expression Σ Γ t → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ t ⟧ₜ + exprs : All (Expression Σ Γ) ts → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ ts ⟧ₜ′ + ref : Reference Σ Γ t → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ t ⟧ₜ + locRef : LocalReference Σ Γ t → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ t ⟧ₜ + assign : Reference Σ Γ t → ⟦ t ⟧ₜ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ + locAssign : LocalReference Σ Γ t → ⟦ t ⟧ₜ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ + stmt : Statement Σ Γ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ + locStmt : LocalStatement Σ Γ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ + fun : Function Σ Γ t → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ t ⟧ₜ + proc : Procedure Σ Γ → ⟦ Σ ⟧ₜ′ × ⟦ Γ ⟧ₜ′ → ⟦ Σ ⟧ₜ′ + + expr (lit {t = t} x) = const (Κ[ t ] x) + expr {Σ = Σ} (state i) = fetch i Σ ∘ proj₁ + expr {Γ = Γ} (var i) = fetch i Γ ∘ proj₂ + expr (e ≟ e₁) = lift ∘ does ∘ uncurry ≈-dec ∘ < expr e , expr e₁ > + expr (e <? e₁) = lift ∘ does ∘ uncurry <-dec ∘ < expr e , expr e₁ > + expr (inv e) = lift ∘ Bool.not ∘ lower ∘ expr e + expr (e && e₁) = lift ∘ uncurry (Bool._∧_ on lower) ∘ < expr e , expr e₁ > + expr (e || e₁) = lift ∘ uncurry (Bool._∨_ on lower) ∘ < expr e , expr e₁ > + expr (not e) = map (lift ∘ 𝔹.¬_ ∘ lower) ∘ expr e + expr (e and e₁) = uncurry (zipWith (lift ∘₂ 𝔹._∧_ on lower)) ∘ < expr e , expr e₁ > + expr (e or e₁) = uncurry (zipWith (lift ∘₂ 𝔹._∨_ on lower)) ∘ < expr e , expr e₁ > + expr [ e ] = (_∷ []) ∘ expr e + expr (unbox e) = Vec.head ∘ expr e + expr (merge e e₁ e₂) = uncurry (uncurry mergeVec) ∘ < < expr e , expr e₁ > , lower ∘ expr e₂ > + expr (slice e e₁) = uncurry sliceVec ∘ < expr e , lower ∘ expr e₁ > + expr (cut e e₁) = uncurry cutVec ∘ < expr e , lower ∘ expr e₁ > + expr (cast eq e) = castVec eq ∘ expr e + expr (- e) = neg ∘ expr e + expr (e + e₁) = uncurry add ∘ < expr e , expr e₁ > + expr (e * e₁) = uncurry mul ∘ < expr e , expr e₁ > + expr (e ^ x) = flip pow x ∘ expr e + expr (e >> n) = lift ∘ flip (shift 2≉0) n ∘ lower ∘ expr e + expr (rnd e) = lift ∘ ⌊_⌋ ∘ lower ∘ expr e + expr (fin {ms = ms} f e) = lift ∘ f ∘ lowerFin ms ∘ expr e + expr (asInt e) = lift ∘ 𝕀⇒ℤ ∘ 𝕀.+_ ∘ Fin.toℕ ∘ lower ∘ expr e + expr nil = const _ + expr (cons {ts = ts} e e₁) = uncurry (cons′ ts) ∘ < expr e , expr e₁ > + expr (head {ts = ts} e) = head′ ts ∘ expr e + expr (tail {ts = ts} e) = tail′ ts ∘ expr e + expr (call f es) = fun f ∘ < proj₁ , exprs es > + expr (if e then e₁ else e₂) = uncurry (uncurry Bool.if_then_else_) ∘ < < lower ∘ expr e , expr e₁ > , expr e₂ > + + exprs [] = const _ + exprs (e ∷ []) = expr e + exprs (e ∷ e₁ ∷ es) = < expr e , exprs (e₁ ∷ es) > + + ref {Σ = Σ} (state i) = fetch i Σ ∘ proj₁ + ref {Γ = Γ} (var i) = fetch i Γ ∘ proj₂ + ref [ r ] = (_∷ []) ∘ ref r + ref (unbox r) = Vec.head ∘ ref r + ref (merge r r₁ e) = uncurry (uncurry mergeVec) ∘ < < ref r , ref r₁ > , lower ∘ expr e > + ref (slice r e) = uncurry sliceVec ∘ < ref r , lower ∘ expr e > + ref (cut r e) = uncurry cutVec ∘ < ref r , lower ∘ expr e > + ref (cast eq r) = castVec eq ∘ ref r + ref nil = const _ + ref (cons {ts = ts} r r₁) = uncurry (cons′ ts) ∘ < ref r , ref r₁ > + ref (head {ts = ts} r) = head′ ts ∘ ref r + ref (tail {ts = ts} r) = tail′ ts ∘ ref r + + locRef {Γ = Γ} (var i) = fetch i Γ ∘ proj₂ + locRef [ r ] = (_∷ []) ∘ locRef r + locRef (unbox r) = Vec.head ∘ locRef r + locRef (merge r r₁ e) = uncurry (uncurry mergeVec) ∘ < < locRef r , locRef r₁ > , lower ∘ expr e > + locRef (slice r e) = uncurry sliceVec ∘ < locRef r , lower ∘ expr e > + locRef (cut r e) = uncurry cutVec ∘ < locRef r , lower ∘ expr e > + locRef (cast eq r) = castVec eq ∘ locRef r + locRef nil = const _ + locRef (cons {ts = ts} r r₁) = uncurry (cons′ ts) ∘ < locRef r , locRef r₁ > + locRef (head {ts = ts} r) = head′ ts ∘ locRef r + locRef (tail {ts = ts} r) = tail′ ts ∘ locRef r + + assign {Σ = Σ} (state i) val σ,γ = < updateAt i Σ val ∘ proj₁ , proj₂ > + assign {Γ = Γ} (var i) val σ,γ = < proj₁ , updateAt i Γ val ∘ proj₂ > + assign [ r ] val σ,γ = assign r (Vec.head val) σ,γ + assign (unbox r) val σ,γ = assign r (val ∷ []) σ,γ + assign (merge r r₁ e) val σ,γ = assign r₁ (cutVec val (lower (expr e σ,γ))) σ,γ ∘ assign r (sliceVec val (lower (expr e σ,γ))) σ,γ + assign (slice r e) val σ,γ = assign r (mergeVec val (cutVec (ref r σ,γ) (lower (expr e σ,γ))) (lower (expr e σ,γ))) σ,γ + assign (cut r e) val σ,γ = assign r (mergeVec (sliceVec (ref r σ,γ) (lower (expr e σ,γ))) val (lower (expr e σ,γ))) σ,γ + assign (cast eq r) val σ,γ = assign r (castVec (sym eq) val) σ,γ + assign nil val σ,γ = id + assign (cons {ts = ts} r r₁) val σ,γ = assign r₁ (tail′ ts val) σ,γ ∘ assign r (head′ ts val) σ,γ + assign (head {ts = ts} r) val σ,γ = assign r (cons′ ts val (ref (tail r) σ,γ)) σ,γ + assign (tail {ts = ts} r) val σ,γ = assign r (cons′ ts (ref (head r) σ,γ) val) σ,γ + + locAssign {Γ = Γ} (var i) val σ,γ = updateAt i Γ val ∘ proj₂ + locAssign [ r ] val σ,γ = locAssign r (Vec.head val) σ,γ + locAssign (unbox r) val σ,γ = locAssign r (val ∷ []) σ,γ + locAssign (merge r r₁ e) val σ,γ = locAssign r₁ (cutVec val (lower (expr e σ,γ))) σ,γ ∘ < proj₁ , locAssign r (sliceVec val (lower (expr e σ,γ))) σ,γ > + locAssign (slice r e) val σ,γ = locAssign r (mergeVec val (cutVec (locRef r σ,γ) (lower (expr e σ,γ))) (lower (expr e σ,γ))) σ,γ + locAssign (cut r e) val σ,γ = locAssign r (mergeVec (sliceVec (locRef r σ,γ) (lower (expr e σ,γ))) val (lower (expr e σ,γ))) σ,γ + locAssign (cast eq r) val σ,γ = locAssign r (castVec (sym eq) val) σ,γ + locAssign nil val σ,γ = proj₂ + locAssign (cons {ts = ts} r r₁) val σ,γ = locAssign r₁ (tail′ ts val) σ,γ ∘ < proj₁ , locAssign r (head′ ts val) σ,γ > + locAssign (head {ts = ts} r) val σ,γ = locAssign r (cons′ ts val (locRef (tail r) σ,γ)) σ,γ + locAssign (tail {ts = ts} r) val σ,γ = locAssign r (cons′ ts (locRef (head r) σ,γ) val) σ,γ + + stmt (s ∙ s₁) = stmt s₁ ∘ stmt s + stmt skip = id + stmt (ref ≔ val) = uncurry (uncurry (assign ref)) ∘ < < expr val , id > , id > + stmt {Γ = Γ} (declare e s) = < proj₁ , tail′ Γ ∘ proj₂ > ∘ stmt s ∘ < proj₁ , uncurry (cons′ Γ) ∘ < expr e , proj₂ > > + stmt (invoke p es) = < proc p ∘ < proj₁ , exprs es > , proj₂ > + stmt (if e then s) = uncurry (uncurry Bool.if_then_else_) ∘ < < lower ∘ expr e , stmt s > , id > + stmt (if e then s else s₁) = uncurry (uncurry Bool.if_then_else_) ∘ < < lower ∘ expr e , stmt s > , stmt s₁ > + stmt {Γ = Γ} (for m s) = Vec.foldl _ (flip λ i → (< proj₁ , tail′ Γ ∘ proj₂ > ∘ stmt s ∘ < proj₁ , cons′ Γ (lift i) ∘ proj₂ >) ∘_) id (Vec.allFin m) + + locStmt (s ∙ s₁) = locStmt s₁ ∘ < proj₁ , locStmt s > + locStmt skip = proj₂ + locStmt (ref ≔ val) = uncurry (uncurry (locAssign ref)) ∘ < < expr val , id > , id > + locStmt {Γ = Γ} (declare e s) = tail′ Γ ∘ locStmt s ∘ < proj₁ , uncurry (cons′ Γ) ∘ < expr e , proj₂ > > + locStmt (if e then s) = uncurry (uncurry Bool.if_then_else_) ∘ < < lower ∘ expr e , locStmt s > , proj₂ > + locStmt (if e then s else s₁) = uncurry (uncurry Bool.if_then_else_) ∘ < < lower ∘ expr e , locStmt s > , locStmt s₁ > + locStmt {Γ = Γ} (for m s) = proj₂ ∘ Vec.foldl _ (flip λ i → (< proj₁ , tail′ Γ ∘ locStmt s > ∘ < proj₁ , cons′ Γ (lift i) ∘ proj₂ >) ∘_) id (Vec.allFin m) + + fun {Γ = Γ} (declare e f) = fun f ∘ < proj₁ , uncurry (cons′ Γ) ∘ < expr e , proj₂ > > + fun (s ∙return e) = expr e ∘ < proj₁ , locStmt s > + + proc (s ∙end) = proj₁ ∘ stmt s |