From 00a0ce9082b4cc1389815defcc806efd4a9b80f4 Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Mon, 18 Apr 2022 15:05:24 +0100 Subject: 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. --- src/Helium/Data/Pseudocode/Manipulate.agda | 2756 +++++++++++++--------------- 1 file changed, 1239 insertions(+), 1517 deletions(-) (limited to 'src/Helium/Data/Pseudocode/Manipulate.agda') 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 (_> 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 (_> 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 (_> 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 (_> 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 > 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 > 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 > 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 > 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 > 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 > 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 > 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 > 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 > 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 > 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 > 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.> 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.> 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 > 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 > ∘ 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) ℕₚ.> 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⇒e0 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> 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 > 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 = ℕₚ.m0 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) ℕₚ.> 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)) -- cgit v1.2.3