├── abt.agda-lib ├── isabelle ├── junk │ ├── Substitution.thy │ └── LambdaExample.thy └── SubstHelper.thy └── src ├── experimental ├── README.md ├── Structures.agda ├── ABT.agda ├── Arith.agda ├── ABTPredicate.agda ├── MapPreserve.agda ├── Fold.agda ├── GenericSubstitution.agda └── AbstractBindingTree.agda ├── Syntax.agda ├── rewriting ├── README.md ├── examples │ ├── Denot.agda │ ├── Gamma.agda │ ├── IfAndOnlyIf.agda │ ├── EquivalenceRelation.agda │ ├── junk │ │ ├── CastSafeLogic.agda │ │ ├── StepIndexedLogic3.agda │ │ └── CastLogRelLemmas.agda │ ├── notes.md │ ├── SystemFUnified.agda │ ├── CastBigStepLogic.agda │ ├── ClosureConversion.agda │ ├── SystemF.agda │ ├── CastDiverge.agda │ └── Experiments.agda └── ABTPredicate.agda ├── Sig.agda ├── Makefile ├── ListAux.agda ├── sub-normal-form ├── Environment.agda ├── Syntax.agda ├── MapPreserve.agda ├── AbstractBindingTree.agda ├── MapFusion.agda └── Map.agda ├── Var.agda ├── ABTPredicate.agda ├── SubstPreserve.agda ├── MapPreserve.agda ├── examples ├── Arith.agda ├── ArithTypeSafety.agda └── Lambda.agda ├── weaken └── AbstractBindingTree.agda ├── GSubst.agda ├── GenericSubstitution.agda ├── MapFusion.agda ├── Map3.agda ├── AbstractBindingTree.agda ├── Renaming.agda ├── Fold2.agda ├── Structures.agda ├── Map.agda ├── ScopedTuple.agda └── Fold.agda /abt.agda-lib: -------------------------------------------------------------------------------- 1 | name: abt 2 | depend: standard-library 3 | include: src 4 | -------------------------------------------------------------------------------- /isabelle/junk/Substitution.thy: -------------------------------------------------------------------------------- 1 | theory Substitution 2 | imports Main AbstractBindingTrees 3 | begin 4 | 5 | 6 | 7 | end -------------------------------------------------------------------------------- /src/experimental/README.md: -------------------------------------------------------------------------------- 1 | 2 | This directory contains experiments in generic functions and theorems 3 | over abstract binding trees. It draws inspiration from the paper "A 4 | Type and Scope Safe Universe of Syntaxes with Binding", ICFP 2018. 5 | -------------------------------------------------------------------------------- /src/Syntax.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | open import Data.List using (List) 3 | open import Data.Nat using (ℕ) 4 | 5 | module Syntax where 6 | 7 | open import Sig public 8 | open import Var public 9 | open import Substitution public 10 | 11 | module OpSig (Op : Set) (sig : Op → List Sig) where 12 | 13 | open import AbstractBindingTree Op sig public 14 | open Substitution.ABTOps Op sig public 15 | open import WellScoped Op sig public 16 | 17 | -------------------------------------------------------------------------------- /src/rewriting/README.md: -------------------------------------------------------------------------------- 1 | 2 | This directory contains an experimental version of the Abstract 3 | Binding Trees library that relies on Agda's rewriting feature to 4 | automate much of the reasoning (i.e. uses the `--rewriting` flag). 5 | The main goal is for all the public-facing definitions to be in terms 6 | of the σ-calculus of Abadi et. al. and to provide all the equations of 7 | the σ-calculus as automatic rewrite rules. Because the σ-calculus is 8 | confluent and strongly normalizing, this automation implements a 9 | decision procedure for equality of substitutions. 10 | -------------------------------------------------------------------------------- /src/Sig.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Data.Nat using (ℕ; zero; suc) 3 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 4 | 5 | module Sig where 6 | 7 | data Sig : Set where 8 | ■ : Sig 9 | ν : Sig → Sig 10 | ∁ : Sig → Sig 11 | 12 | sig→ℕ : Sig → ℕ 13 | sig→ℕ ■ = 0 14 | sig→ℕ (ν b) = suc (sig→ℕ b) 15 | sig→ℕ (∁ b) = sig→ℕ b 16 | 17 | ℕ→sig : ℕ → Sig 18 | ℕ→sig 0 = ■ 19 | ℕ→sig (suc n) = ν (ℕ→sig n) 20 | 21 | {- The following is used in Fold2 -} 22 | 23 | Result : {ℓ : Level} → Set ℓ → Sig → Set ℓ 24 | Result V ■ = V 25 | Result V (ν b) = V → Result V b 26 | Result V (∁ b) = Result V b 27 | 28 | -------------------------------------------------------------------------------- /src/Makefile: -------------------------------------------------------------------------------- 1 | ROOT = /Users/jsiek/abstract-binding-trees 2 | BUILD_DIR = $(ROOT)/_build/2.6.1/agda 3 | 4 | AGDA = ScopedTuple.agda \ 5 | Var.agda \ 6 | AbstractBindingTree.agda \ 7 | Map.agda \ 8 | GSubst.agda \ 9 | GenericSubstitution.agda \ 10 | Renaming.agda \ 11 | MapFusion.agda \ 12 | Substitution.agda \ 13 | ABTPredicate.agda \ 14 | Fold.agda \ 15 | MapPreserve.agda \ 16 | WellScoped.agda \ 17 | FoldPreserve.agda \ 18 | FoldMapFusion.agda \ 19 | FoldFoldFusion.agda \ 20 | Syntax.agda \ 21 | examples/Lambda.agda \ 22 | examples/Arith.agda \ 23 | examples/BlogTypeSafetyTwoEasy.lagda.md 24 | 25 | AGDAI = $(AGDA:%.agda=$(BUILD_DIR)/src/%.agdai) 26 | 27 | $(BUILD_DIR)/src/%.agdai: %.agda 28 | agda $< 29 | 30 | all: $(AGDA) $(AGDAI) 31 | 32 | clean: 33 | rm -f $(BUILD_DIR)/src/*.agdai $(BUILD_DIR)/src/examples/*.agdai *~ 34 | -------------------------------------------------------------------------------- /src/ListAux.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Data.List using (List; []; _∷_; length; _++_) 3 | open import Data.List.Properties using (length-++; ++-assoc; ≡-dec) 4 | open import Data.Nat using (ℕ; zero; suc; _+_; _<_; _≤_; z≤n; s≤s; _≟_) 5 | open import Data.Nat.Properties 6 | using (+-suc; ≤-trans; ≤-step; ≤-refl; ≤-reflexive; m≤m+n; ≤-pred) 7 | import Relation.Binary.PropositionalEquality as Eq 8 | open Eq using (_≡_; refl; sym; trans; cong; cong₂; cong-app; subst) 9 | open Eq.≡-Reasoning 10 | 11 | module ListAux where 12 | 13 | private 14 | variable I : Set 15 | 16 | nth : (xs : List I) → (i : ℕ) → .{i< : i < length xs} → I 17 | nth (x ∷ xs) zero {i<} = x 18 | nth (x ∷ xs) (suc i) {i<} = nth xs i {≤-pred i<} 19 | 20 | nth-cong : ∀ (xs ys : List I) (i : ℕ) 21 | {p : i < length xs }{q : i < length ys } 22 | → xs ≡ ys 23 | → nth xs i {p} ≡ nth ys i {q} 24 | nth-cong xs ys i {p}{q} refl = refl 25 | 26 | length-++-< : (xs ys : List I) (y : I) 27 | → length xs < length (xs ++ y ∷ ys) 28 | length-++-< xs ys y rewrite length-++ xs {y ∷ ys} 29 | | +-suc (length xs) (length ys) = s≤s (m≤m+n _ _) 30 | 31 | nth-++ : ∀ (xs ys : List I) (y : I) 32 | → nth (xs ++ (y ∷ ys)) (length xs) {length-++-< xs ys y} ≡ y 33 | nth-++ [] ys y = refl 34 | nth-++ (x ∷ xs) ys y = nth-++ xs ys y 35 | -------------------------------------------------------------------------------- /src/rewriting/examples/Denot.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | {- 3 | UNDER CONSTRUCTION 4 | -} 5 | 6 | open import Data.Empty using (⊥; ⊥-elim) 7 | open import Data.List.Membership.Propositional 8 | open import Data.List using (List; []; _∷_; length) 9 | open import Data.Nat using (ℕ; zero; suc) 10 | open import Data.Product using (_×_; _,_; proj₁; proj₂; ∃-syntax; Σ-syntax) 11 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym) 12 | open import Var 13 | 14 | module rewriting.examples.Denot where 15 | 16 | data Val : Set where 17 | lit : ℕ → Val 18 | _↦_ : List Val → Val → Val 19 | pair : List Val → List Val → Val 20 | 21 | D : Set₁ 22 | D = Val → Set 23 | 24 | ⊥′ : D 25 | ⊥′ v = ⊥ 26 | 27 | {- Application Operator -} 28 | infixl 7 _●_ 29 | _●_ : D → D → D 30 | _●_ D₁ D₂ w = (∃[ v ] D₂ v) 31 | × Σ[ V ∈ List Val ] D₁ (V ↦ w) × (∀ {v} → v ∈ V → D₂ v) 32 | 33 | {- Abstraction Operator -} 34 | Λ : (D → D) → D 35 | Λ F (lit k) = ⊥ 36 | Λ F (V ↦ w) = F (_∈ V) w 37 | Λ F (pair v w) = ⊥ 38 | 39 | {- Pair operator -} 40 | ⦅_,_⦆ : D → D → D 41 | ⦅ D₁ , D₂ ⦆ (lit k) = ⊥ 42 | ⦅ D₁ , D₂ ⦆ (v ↦ w) = ⊥ 43 | ⦅ D₁ , D₂ ⦆ (pair V W) = (∀ {v} → v ∈ V → D₁ v) × (∀ {w} → w ∈ W → D₂ w) 44 | 45 | {- Fst operator -} 46 | π₁ : D → D 47 | π₁ D v = ∃[ V ] ∃[ W ] D (pair V W) × v ∈ V 48 | 49 | {- Snd operator -} 50 | π₂ : D → D 51 | π₂ D w = ∃[ V ] ∃[ W ] D (pair V W) × w ∈ W 52 | 53 | {- For looking up variables -} 54 | nth : ∀{ℓ}{A : Set ℓ} → ℕ → List A → A → A 55 | nth n [] d = d 56 | nth zero (x ∷ xs) d = x 57 | nth (suc n) (x ∷ xs) d = nth n xs d 58 | 59 | -------------------------------------------------------------------------------- /src/sub-normal-form/Environment.agda: -------------------------------------------------------------------------------- 1 | open import Data.List using (List; []; _∷_) 2 | open import Data.Nat using (ℕ; zero; suc; _≟_) 3 | open import Function using (_∘_) 4 | import Relation.Binary.PropositionalEquality as Eq 5 | open Eq using (_≡_; refl; sym; cong; cong₂; cong-app) 6 | open import Relation.Nullary using (Dec; yes; no) 7 | open import Var 8 | 9 | module Environment where 10 | 11 | record Shiftable {ℓ} (V : Set ℓ) : Set ℓ where 12 | field ⇑ : V → V 13 | var→val : Var → V 14 | var→val-suc-shift : ∀{x} → var→val (suc x) ≡ ⇑ (var→val x) 15 | 16 | instance 17 | Var-is-Shiftable : Shiftable Var 18 | Var-is-Shiftable = record { var→val = λ x → x ; ⇑ = suc 19 | ; var→val-suc-shift = λ {x} → refl } 20 | 21 | open Shiftable {{...}} 22 | 23 | l→f : List Var → Var → Var 24 | l→f [] y = y 25 | l→f (x ∷ xs) zero = x 26 | l→f (x ∷ xs) (suc y) = l→f xs y 27 | 28 | record Env {ℓ} (E : Set ℓ) (V : Set ℓ) : Set ℓ where 29 | field {{V-is-Shiftable}} : Shiftable V 30 | field ⟅_⟆ : E → Var → V {- lookup variable in environment -} 31 | _,_ : E → V → E {- update environment, mapping 0 to value -} 32 | ⟰ : E → E {- shift up by one everything in environment -} 33 | lookup-0 : ∀ ρ v → ⟅ ρ , v ⟆ 0 ≡ v 34 | lookup-suc : ∀ ρ v x → ⟅ ρ , v ⟆ (suc x) ≡ ⇑ (⟅ ρ ⟆ x) 35 | lookup-shift : ∀ ρ x → ⟅ ⟰ ρ ⟆ x ≡ ⇑ (⟅ ρ ⟆ x) 36 | ext : E → E 37 | ext ρ = ρ , (var→val 0) 38 | 39 | fun-extend : ∀{ℓ}{V : Set ℓ}{{_ : Shiftable V}} → (Var → V) → V → (Var → V) 40 | fun-extend ρ v zero = v 41 | fun-extend ρ v (suc x) = ⇑ (ρ x) 42 | 43 | instance 44 | Fun-is-Env : ∀{ℓ}{V : Set ℓ}{{_ : Shiftable V}} → Env (Var → V) V 45 | Fun-is-Env = record { ⟅_⟆ = λ ρ x → ρ x ; _,_ = fun-extend 46 | ; ⟰ = λ ρ x → ⇑ (ρ x) ; lookup-0 = λ ρ v → refl 47 | ; lookup-suc = λ ρ v x → refl ; lookup-shift = λ ρ x → refl } 48 | -------------------------------------------------------------------------------- /src/rewriting/examples/Gamma.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | {- 3 | Intermediate language prior to key step in closure conversion. 4 | -} 5 | 6 | open import Data.List using (List; []; _∷_; length) 7 | open import Data.Nat using (ℕ; zero; suc) 8 | open import Data.Product using (_×_; _,_; proj₁; proj₂; ∃-syntax) 9 | open import Data.Unit.Polymorphic using (⊤; tt) 10 | open import Data.Vec using (Vec) renaming ([] to []̌; _∷_ to _∷̌_) 11 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym) 12 | open import Var 13 | open import Sig 14 | open import rewriting.examples.Denot 15 | 16 | module rewriting.examples.Gamma where 17 | 18 | data Op : Set where 19 | op-gamma : Op 20 | op-app : Op 21 | op-lit : ℕ → Op 22 | op-cons : Op 23 | op-fst : Op 24 | op-snd : Op 25 | 26 | sig : Op → List Sig 27 | sig op-gamma = ν (ν ■) ∷ ■ ∷ [] 28 | sig op-app = ■ ∷ ■ ∷ [] 29 | sig (op-lit k) = [] 30 | sig op-cons = ■ ∷ ■ ∷ [] 31 | sig op-fst = ■ ∷ [] 32 | sig op-snd = ■ ∷ [] 33 | 34 | open import rewriting.AbstractBindingTree Op sig hiding (`_) 35 | open import rewriting.AbstractBindingTree Op sig 36 | using (`_) renaming (ABT to Term) public 37 | 38 | pattern $ k = op-lit k ⦅ nil ⦆ 39 | 40 | pattern γ N M = op-gamma ⦅ cons (bind (bind (ast N))) (cons (ast M) nil) ⦆ 41 | 42 | infixl 7 _·_ 43 | pattern _·_ L M = op-app ⦅ cons (ast L) (cons (ast M) nil) ⦆ 44 | 45 | infixl 7 ⟨_,_⟩ 46 | pattern ⟨_,_⟩ L M = op-cons ⦅ cons (ast L) (cons (ast M) nil) ⦆ 47 | 48 | pattern fst L = op-fst ⦅ (cons (ast L) nil) ⦆ 49 | 50 | pattern snd L = op-snd ⦅ (cons (ast L) nil) ⦆ 51 | 52 | ⟦_⟧ : Term → (List D) → D 53 | ⟦ ` x ⟧ ρ = nth x ρ ⊥′ 54 | ⟦ $ k ⟧ ρ v = (v ≡ lit k) 55 | ⟦ γ N M ⟧ ρ w = (∃[ v ] ⟦ M ⟧ ρ v) × (Λ (λ D → ⟦ N ⟧ (⟦ M ⟧ ρ ∷ D ∷ [])) w) 56 | ⟦ L · M ⟧ ρ = ⟦ L ⟧ ρ ● ⟦ M ⟧ ρ 57 | ⟦ ⟨ L , M ⟩ ⟧ ρ = ⦅ ⟦ L ⟧ ρ , ⟦ M ⟧ ρ ⦆ 58 | ⟦ fst L ⟧ ρ = π₁ (⟦ L ⟧ ρ) 59 | ⟦ snd L ⟧ ρ = π₂ (⟦ L ⟧ ρ) 60 | 61 | 62 | -------------------------------------------------------------------------------- /src/Var.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 3 | open import Data.Nat using (ℕ; zero; suc; _<_; z≤n; s≤s; _+_) 4 | open import Data.Empty.Polymorphic using (⊥) 5 | open import Data.Unit.Polymorphic using (⊤) 6 | open import Data.List using (List; []; _∷_; length; _++_) 7 | open import Data.Product using (_×_) 8 | open import Level using (levelOfType) 9 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 10 | open import Sig 11 | 12 | module Var where 13 | 14 | Var : Set 15 | Var = ℕ 16 | 17 | {- 18 | data Lift (ℓᵛ : Level) {ℓᶜ : Level} (C : Set ℓᶜ) : Set (ℓᵛ ⊔ ℓᶜ) where 19 | lift : C → Lift ℓᵛ C 20 | 21 | lower : ∀{ℓᵛ ℓᶜ}{C : Set ℓᶜ} → Lift ℓᵛ C → C 22 | lower (lift c) = c 23 | 24 | lift-lower-id : ∀{ℓᵛ ℓᶜ}{C : Set ℓᶜ} (lc : Lift ℓᵛ C) 25 | → lift (lower lc) ≡ lc 26 | lift-lower-id (lift c) = refl 27 | -} 28 | 29 | private 30 | variable 31 | ℓ : Level 32 | I : Set ℓ 33 | Γ Δ : List I 34 | x : Var 35 | A : I 36 | 37 | _∋_⦂_ : List I → Var → I → Set (levelOfType I) 38 | _∋_⦂_ {I = I} [] x A = ⊥ 39 | _∋_⦂_ (B ∷ Γ) zero A = A ≡ B 40 | _∋_⦂_ (B ∷ Γ) (suc x) A = Γ ∋ x ⦂ A 41 | 42 | ∋x→< : Γ ∋ x ⦂ A → x < (length Γ) 43 | ∋x→< {Γ = []} {x} () 44 | ∋x→< {Γ = B ∷ Γ} {x = zero} ∋x = s≤s z≤n 45 | ∋x→< {Γ = B ∷ Γ} {x = suc x} ∋x = s≤s (∋x→< {Γ = Γ} ∋x) 46 | 47 | <→∋x : ∀{Γ : List {a = lzero} ⊤}{x A} → x < (length Γ) → Γ ∋ x ⦂ A 48 | <→∋x {Γ = B ∷ Γ} {x = zero} x<Γ = refl 49 | <→∋x {Γ = B ∷ Γ} {x = suc x} (s≤s x<Γ) = <→∋x {Γ = Γ} x<Γ 50 | 51 | ∋++ : Γ ∋ x ⦂ A → (Δ ++ Γ) ∋ (length Δ + x) ⦂ A 52 | ∋++ {Δ = []} ∋ΔΓ = ∋ΔΓ 53 | ∋++ {Δ = B ∷ Δ} ∋ΔΓ = ∋++ {Δ = Δ} ∋ΔΓ 54 | 55 | {--- types for bound variables ---} 56 | 57 | BType : ∀{ℓ} → Set ℓ → Sig → Set ℓ 58 | BType I ■ = ⊤ 59 | BType I (ν b) = I × BType I b 60 | BType I (∁ b) = BType I b 61 | 62 | BTypes : ∀{ℓ} → Set ℓ → List Sig → Set ℓ 63 | BTypes I [] = ⊤ 64 | BTypes I (b ∷ bs) = BType I b × BTypes I bs 65 | 66 | -------------------------------------------------------------------------------- /src/rewriting/examples/IfAndOnlyIf.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | 3 | module rewriting.examples.IfAndOnlyIf where 4 | 5 | open import Data.Product 6 | using (_×_; _,_; proj₁; proj₂; Σ; ∃; Σ-syntax; ∃-syntax) 7 | import Relation.Binary.PropositionalEquality as Eq 8 | open Eq using (_≡_; refl; sym; trans; cong; cong₂; cong-app; subst) 9 | 10 | abstract 11 | _⇔_ : Set → Set → Set 12 | P ⇔ Q = (P → Q) × (Q → P) 13 | 14 | ⇔-intro : ∀{P Q : Set} 15 | → (P → Q) 16 | → (Q → P) 17 | ------- 18 | → P ⇔ Q 19 | ⇔-intro PQ QP = PQ , QP 20 | 21 | ⇔-elim : ∀{P Q : Set} 22 | → P ⇔ Q 23 | ----------------- 24 | → (P → Q) × (Q → P) 25 | ⇔-elim PQ = PQ 26 | 27 | ⇔-to : ∀{P Q : Set} 28 | → P ⇔ Q 29 | ------- 30 | → (P → Q) 31 | ⇔-to PQ = proj₁ PQ 32 | 33 | ⇔-fro : ∀{P Q : Set} 34 | → P ⇔ Q 35 | ------- 36 | → (Q → P) 37 | ⇔-fro PQ = proj₂ PQ 38 | 39 | ⇔-refl : ∀{P Q : Set} 40 | → P ≡ Q 41 | ------ 42 | → P ⇔ Q 43 | ⇔-refl refl = (λ x → x) , (λ x → x) 44 | 45 | ⇔-sym : ∀{P Q : Set} 46 | → P ⇔ Q 47 | ------ 48 | → Q ⇔ P 49 | ⇔-sym PQ = (proj₂ PQ) , (proj₁ PQ) 50 | 51 | ⇔-trans : ∀{P Q R : Set} 52 | → P ⇔ Q 53 | → Q ⇔ R 54 | ------ 55 | → P ⇔ R 56 | ⇔-trans PQ QR = 57 | (λ z → proj₁ QR (proj₁ PQ z)) , (λ z → proj₂ PQ (proj₂ QR z)) 58 | 59 | infixr 2 _⇔⟨_⟩_ 60 | infix 3 _QED 61 | 62 | _⇔⟨_⟩_ : 63 | (P : Set) 64 | → ∀{Q} → P ⇔ Q 65 | → ∀{R} → Q ⇔ R 66 | ------------- 67 | → P ⇔ R 68 | P ⇔⟨ PQ ⟩ QR = ⇔-trans PQ QR 69 | 70 | _QED : 71 | (P : Set) 72 | --------- 73 | → P ⇔ P 74 | P QED = ⇔-refl refl 75 | 76 | abstract 77 | ×-cong-⇔ : ∀{S S′ T T′} 78 | → S ⇔ S′ 79 | → T ⇔ T′ 80 | ------------------- 81 | → (S × T) ⇔ (S′ × T′) 82 | ×-cong-⇔ SS′ TT′ = (λ x → (proj₁ SS′ (proj₁ x)) , (proj₁ TT′ (proj₂ x))) 83 | , (λ x → (proj₂ SS′ (proj₁ x)) , (proj₂ TT′ (proj₂ x))) 84 | -------------------------------------------------------------------------------- /src/experimental/Structures.agda: -------------------------------------------------------------------------------- 1 | open import Agda.Primitive using (Level; lzero; lsuc) 2 | open import Data.Nat using (ℕ; zero; suc) 3 | import Relation.Binary.PropositionalEquality as Eq 4 | open Eq using (_≡_; refl; sym; cong; cong₂; cong-app) 5 | open import Var 6 | 7 | module Structures where 8 | 9 | record Shiftable {ℓ} (V : Set ℓ) : Set ℓ where 10 | field ⇑ : V → V 11 | var→val : Var → V 12 | var→val-suc-shift : ∀{x} → var→val (suc x) ≡ ⇑ (var→val x) 13 | 14 | open Shiftable {{...}} public 15 | 16 | instance 17 | Var-is-Shiftable : Shiftable Var 18 | Var-is-Shiftable = record { var→val = λ x → x ; ⇑ = suc 19 | ; var→val-suc-shift = λ {x} → refl } 20 | 21 | 22 | record Composable {ℓ} (V₁ V₂ V₃ : Set ℓ){{_ : Shiftable V₁}} : Set ℓ where 23 | field ⌈_⌉ : (Var → V₂) → V₁ → V₃ 24 | val₂₃ : V₂ → V₃ 25 | ⌈⌉-var→val : ∀ σ x → ⌈ σ ⌉ (var→val x) ≡ val₂₃ (σ x) 26 | 27 | open Composable {{...}} public 28 | 29 | instance 30 | Var³-Composable : Composable Var Var Var 31 | Var³-Composable = record { ⌈_⌉ = λ f x → f x ; val₂₃ = λ x → x 32 | ; ⌈⌉-var→val = λ σ x → refl } 33 | 34 | infixr 5 _⨟_ 35 | 36 | _⨟_ : ∀{ℓ}{V₁ V₂ V₃ : Set ℓ} {{_ : Shiftable V₁}} {{_ : Composable V₁ V₂ V₃}} 37 | → (Var → V₁) → (Var → V₂) → (Var → V₃) 38 | (σ₁ ⨟ σ₂) x = ⌈ σ₂ ⌉ (σ₁ x) 39 | 40 | record Relatable {ℓ} (V₁ V₂ : Set ℓ) 41 | {{S₁ : Shiftable V₁}}{{S₂ : Shiftable V₂}} : Set (lsuc ℓ) where 42 | field _∼_ : V₁ → V₂ → Set 43 | var→val∼ : ∀ x → var→val{ℓ}{V₁} x ∼ var→val{ℓ}{V₂} x 44 | shift∼ : ∀{v₁ v₂}→ v₁ ∼ v₂ → ⇑ v₁ ∼ ⇑ v₂ 45 | 46 | open Relatable {{...}} public 47 | 48 | record Renameable {ℓ} (V : Set ℓ) : Set ℓ where 49 | field ren : (Var → Var) → V → V 50 | 51 | open Renameable {{...}} public 52 | 53 | instance 54 | Var-Renameable : Renameable Var 55 | Var-Renameable = record { ren = λ f x → f x } 56 | 57 | postulate 58 | extensionality : ∀{ℓ₁ ℓ₂} {A : Set ℓ₁ }{B : Set ℓ₂} {f g : A → B} 59 | → (∀ (x : A) → f x ≡ g x) 60 | ----------------------- 61 | → f ≡ g 62 | 63 | -------------------------------------------------------------------------------- /src/experimental/ABT.agda: -------------------------------------------------------------------------------- 1 | open import Data.List using (List; []; _∷_) 2 | open import Data.Nat using (ℕ; zero; suc; _⊔_) 3 | open import Data.Product using (_×_; proj₁; proj₂) renaming (_,_ to ⟨_,_⟩ ) 4 | open import ScopedTuple using (Sig; Tuple; foldr; map; zip) 5 | import Relation.Binary.PropositionalEquality as Eq 6 | open Eq using (_≡_; refl; sym; cong; cong₂; cong-app) 7 | open import Size 8 | open import Var 9 | 10 | module experimental.ABT (Op : Set) (sig : Op → Sig) where 11 | 12 | data Term : Size → Set where 13 | `_ : ∀{s} → Var → Term (↑ s) 14 | _⦅_⦆ : ∀{s} → (op : Op) → Tuple (sig op) (λ _ → Term s) → Term (↑ s) 15 | 16 | ABT : Set 17 | ABT = Term ∞ 18 | 19 | 20 | {- An example of a recursive function on Terms. -} 21 | 22 | depth : ∀{s} → Term s → ℕ 23 | depth (` x) = 0 24 | depth (op ⦅ args ⦆) = foldr _⊔_ 0 (map depth args) 25 | 26 | {- some basic properties -} 27 | 28 | var-injective : ∀ {x y} → ` x ≡ ` y → x ≡ y 29 | var-injective refl = refl 30 | 31 | {- Heterogeneous equality between terms of different size -} 32 | 33 | data _⩭_ : {s : Size} → Term s → {t : Size} → Term t → Set 34 | ⟨_⟩_⩭_ : ∀{s t : Size} → (bs : Sig) → Tuple bs (λ _ → Term s) 35 | → Tuple bs (λ _ → Term t) → Set 36 | 37 | data _⩭_ where 38 | var⩭ : ∀{i j : Size}{k l : Var} → k ≡ l → `_ {s = i} k ⩭ `_ {s = j} l 39 | node⩭ : ∀{i j : Size}{op}{args args'} 40 | → ⟨ sig op ⟩ args ⩭ args' 41 | → _⦅_⦆ {s = i} op args ⩭ _⦅_⦆ {s = j} op args' 42 | 43 | ⟨ bs ⟩ xs ⩭ ys = zip (λ M N → M ⩭ N) {bs} xs ys 44 | 45 | 46 | {- Conversion from heterogeneous equality to equality -} 47 | 48 | ⩭→≡ : ∀ {s : Size}{M N : Term s} → M ⩭ N → M ≡ N 49 | ⟨_⟩⩭→≡ : ∀{s : Size} → (bs : Sig) → {xs ys : Tuple bs (λ _ → Term s)} 50 | → ⟨ bs ⟩ xs ⩭ ys → xs ≡ ys 51 | 52 | ⩭→≡ {.(Size.↑ _)} {.(` _)} {.(` _)} (var⩭ refl) = refl 53 | ⩭→≡ {.(Size.↑ _)} {.(_ ⦅ _ ⦆)} {.(_ ⦅ _ ⦆)} (node⩭ {op = op} args⩭) = 54 | cong (_⦅_⦆ op) (⟨ sig op ⟩⩭→≡ args⩭) 55 | 56 | ⟨_⟩⩭→≡ {s} [] {tt} {tt} tt = refl 57 | ⟨_⟩⩭→≡ {s} (b ∷ bs) {⟨ x , xs ⟩} {⟨ y , ys ⟩} ⟨ x=y , xs=ys ⟩ 58 | rewrite ⩭→≡ x=y | ⟨ bs ⟩⩭→≡ xs=ys = refl 59 | 60 | -------------------------------------------------------------------------------- /src/sub-normal-form/Syntax.agda: -------------------------------------------------------------------------------- 1 | open import Data.Bool using (Bool; true; false; _∨_) 2 | open import Data.Empty using (⊥; ⊥-elim) 3 | open import Data.Empty.Irrelevant renaming (⊥-elim to ⊥-elimi) 4 | open import Data.List using (List; []; _∷_; length) 5 | open import Data.Nat 6 | using (ℕ; zero; suc; _+_; pred; _≤_; _<_; _≟_; s≤s; z≤n; _≤?_) 7 | open import Data.Nat.Properties 8 | using (+-comm; +-suc; ≤-pred; m≤m+n; ≤⇒≯; suc-injective) 9 | open import Data.Product using (_×_) renaming (_,_ to ⟨_,_⟩) 10 | open import Data.Sum using (_⊎_; inj₁; inj₂) 11 | open import Data.Unit using (⊤; tt) 12 | open import Function using (_∘_) 13 | import GenericSubstitution 14 | import Relation.Binary.PropositionalEquality as Eq 15 | open Eq using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app) 16 | open Eq.≡-Reasoning 17 | open import Relation.Nullary using (¬_; Dec; yes; no) 18 | 19 | module Syntax where 20 | 21 | open import Var public 22 | 23 | open import Substitution public 24 | 25 | module OpSig (Op : Set) (sig : Op → List ℕ) where 26 | 27 | open import AbstractBindingTree Op sig public 28 | open import Environment 29 | open Env {{...}} 30 | open Substitution.ABTOps Op sig public 31 | open import Map using (map; map-arg; map-args) 32 | 33 | open import WellScoped Op sig public 34 | 35 | {---------------------------------------------------------------------------- 36 | Convert (Var → Var) Function to Renaming 37 | ----------------------------------------------------------------------------} 38 | 39 | private 40 | make-ren : (Var → Var) → ℕ → ℕ → Rename 41 | make-ren ρ x zero = ↑ 0 42 | make-ren ρ x (suc m) = ρ x • make-ren ρ (suc x) m 43 | 44 | ⟅make-ren⟆ : ∀{m}{x}{i}{ρ} 45 | → x < m 46 | → ⟅ make-ren ρ i m ⟆ x ≡ ρ (x + i) 47 | ⟅make-ren⟆ {suc m} {zero} {i} {ρ} x>=_ : Maybe Val → (Val → Maybe Val) → Maybe Val 42 | x >>= f 43 | with x 44 | ... | nothing = nothing 45 | ... | just n = f n 46 | 47 | eval-op : (op : Op) → Tuple (sig op) (Bind (Maybe Val) (Maybe Val)) 48 | → Maybe Val 49 | eval-op (op-num n) tt = just (v-num n) 50 | eval-op op-mult ⟨ x , ⟨ y , tt ⟩ ⟩ = do n ← x; m ← y; just (v-num (n * m)) 51 | eval-op op-let ⟨ x , ⟨ f , tt ⟩ ⟩ = do n ← x; f (just n) 52 | eval-op (op-bool b) tt = just (v-bool b) 53 | 54 | S : Substable (Maybe Val) 55 | S = record { var→val = λ x → nothing ; shift = λ r → r 56 | ; var→val-suc-shift = refl } 57 | 58 | E : Fold (Maybe Val) (Maybe Val) 59 | E = record { S = S ; ret = λ x → x ; fold-op = eval-op } 60 | open Fold E 61 | 62 | eval : AST → Maybe Val 63 | eval = fold (↑ 0) 64 | 65 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym) 66 | 67 | _ : eval ($ 2 ⊗ $ 21) ≡ just 42 68 | _ = refl 69 | 70 | _ : eval (` 0) ≡ nothing 71 | _ = refl 72 | 73 | _ : eval (bind $ 21 { $ 2 ⊗ ` 0 }) ≡ just 42 74 | _ = refl 75 | 76 | _ : eval (bind ` 0 { $ 2 ⊗ $ 21 }) ≡ nothing 77 | _ = refl 78 | 79 | -------------------------------------------------------------------------------- /src/ABTPredicate.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 3 | open import Data.List using (List; []; _∷_; length; map; foldl) 4 | open import Data.Nat using (ℕ; zero; suc; _+_; _<_; _≤_; _⊔_; z≤n; s≤s) 5 | open import Data.Nat.Properties 6 | using (+-suc; ≤-trans; ≤-step; ≤-refl; ≤-reflexive; m≤m+n; ≤-pred; 7 | m≤m⊔n; n≤m⊔n; m≤n⇒m Δ 42 | 43 | But is f surjective? If not, then for a variable x in Δ that 44 | is not in the image of f, we don't know that f⁻¹ x is in Γ. 45 | 46 | -} 47 | image-max : (Γ : List I) → (f : Var → Var) 48 | → Σ[ m ∈ ℕ ] (∀ x → x < length Γ → f x ≤ m) 49 | image-max [] f = ⟨ 0 , (λ _ ()) ⟩ 50 | image-max (A ∷ Γ) f 51 | with image-max Γ f 52 | ... | ⟨ m , x<Γ→fx≤m ⟩ = 53 | ⟨ f (length Γ) ⊔ m , G ⟩ 54 | where 55 | G : (x : ℕ) → suc x ≤ suc (length Γ) 56 | → f x ≤ f (length Γ) ⊔ m 57 | G x (s≤s lt) 58 | with m≤n⇒m>=_ : ∀{V : Set} → (Maybe V) → (V → Maybe V) → Maybe V 64 | x >>= f 65 | with x 66 | ... | nothing = nothing 67 | ... | just n = f n 68 | 69 | num? : ∀{V : Set} → Val → (ℕ → Maybe V) → Maybe V 70 | num? mv f 71 | with mv 72 | ... | v-num n = f n 73 | ... | _ = nothing 74 | 75 | bool? : ∀{V : Set} → Val → (𝔹 → Maybe V) → Maybe V 76 | bool? mv f 77 | with mv 78 | ... | v-bool b = f b 79 | ... | _ = nothing 80 | 81 | 82 | eval-op : (op : Op) → Tuple (sig op) (Result (Maybe Val)) 83 | → Maybe Val 84 | eval-op (op-num n) tt = just (v-num n) 85 | eval-op op-error tt = nothing 86 | eval-op op-mult ⟨ x , ⟨ y , tt ⟩ ⟩ = do 87 | v₁ ← x ; v₂ ← y 88 | num? v₁ (λ n → num? v₂ (λ m → just (v-num (n * m)))) 89 | eval-op op-let ⟨ mv , ⟨ f , tt ⟩ ⟩ = f mv 90 | {- skipping check on mv, simpler -} 91 | eval-op (op-bool b) tt = just (v-bool b) 92 | eval-op op-if ⟨ cnd , ⟨ thn , ⟨ els , tt ⟩ ⟩ ⟩ = do 93 | vᶜ ← cnd 94 | bool? vᶜ (λ b → if b then thn else els) 95 | 96 | open Structures.WithOpSig Op sig 97 | 98 | eval : (Var → Maybe Val) → AST → Maybe Val 99 | eval = fold eval-op nothing 100 | 101 | evaluate : AST → Maybe Val 102 | evaluate M = eval (λ x → nothing) M 103 | 104 | _ : evaluate ($ 2 ⊗ $ 21) ≡ just (v-num 42) 105 | _ = refl 106 | 107 | _ : evaluate (` 0) ≡ nothing 108 | _ = refl 109 | 110 | _ : evaluate (bind $ 21 { $ 2 ⊗ ` 0 }) ≡ just (v-num 42) 111 | _ = refl 112 | 113 | _ : evaluate (bind ` 0 { $ 2 ⊗ $ 21 }) ≡ just (v-num 42) 114 | _ = refl {- call-by-name behavior wrt. errors because skipped check -} 115 | 116 | 117 | -------------------------------------------------------------------------------- /src/rewriting/examples/CastBigStepLogic.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module rewriting.examples.CastBigStepLogic where 3 | 4 | open import Data.List using (List; []; _∷_; length; map) 5 | open import Data.Nat 6 | open import Data.Bool using (true; false) renaming (Bool to 𝔹) 7 | open import Data.Nat.Properties 8 | open import Data.Product using (_,_;_×_; proj₁; proj₂; Σ-syntax; ∃-syntax) 9 | open import Data.Unit using (⊤; tt) 10 | open import Data.Unit.Polymorphic using () renaming (⊤ to topᵖ; tt to ttᵖ) 11 | open import Data.Empty using (⊥; ⊥-elim) 12 | open import Data.Sum using (_⊎_; inj₁; inj₂) 13 | open import Relation.Binary.PropositionalEquality as Eq 14 | using (_≡_; _≢_; refl; sym; cong; subst; trans) 15 | open import Relation.Nullary using (¬_; Dec; yes; no) 16 | open import Var 17 | open import rewriting.examples.Cast 18 | open import rewriting.examples.CastBigStepResult 19 | open import rewriting.examples.StepIndexedLogic2 20 | 21 | infixr 6 _⇓_ 22 | _⇓_ : Term → Result → ℕ → Set 23 | (M ⇓ R) 0 = ⊤ 24 | (M ⇓ R) (suc k) = Halt R × ∃[ n ] (M ⇓ᵏ R) (n ∸ (suc k)) 25 | 26 | ⇓-value : ∀ V k → Value V → (V ⇓ val V) k 27 | ⇓-value V zero v = tt 28 | ⇓-value V (suc k) v 29 | with ⇓ᵏ-value V v 30 | ... | l , V⇓Vl = valueH , l + suc k , Goal 31 | where 32 | Goal : (V ⇓ᵏ val V) (l + suc k ∸ suc k) 33 | Goal rewrite +-∸-assoc l {suc k}{suc k} ≤-refl | n∸n≡0 k | +-identityʳ l = 34 | V⇓Vl 35 | 36 | sk∸k=1 : ∀ k → suc k ∸ k ≡ suc zero 37 | sk∸k=1 zero = refl 38 | sk∸k=1 (suc k) = sk∸k=1 k 39 | 40 | ⇓-blame : ∀{k} → (blame ⇓ blameR) k 41 | ⇓-blame {zero} = tt 42 | ⇓-blame {suc k} = blameH , ((suc (suc k)) , Goal) 43 | where 44 | Goal : (blame ⇓ᵏ blameR) (suc k ∸ k) 45 | Goal rewrite sk∸k=1 k = blame⇓ᵏ 46 | 47 | {- false 48 | ⇓-timeout : ∀{M}{k} → (M ⇓ timeout) k 49 | ⇓-timeout {M} {zero} = tt 50 | ⇓-timeout {M} {suc k} = {!!} , {!!} 51 | -} 52 | 53 | downClosed⇓ : ∀ M R → downClosed (M ⇓ R) 54 | downClosed⇓ M R zero M⇓ zero z≤n = tt 55 | downClosed⇓ M R (suc k) (H , n , M⇓Rn-k) zero z≤n = tt 56 | downClosed⇓ M R (suc k) (H , n , M⇓Rn-k) (suc j) (s≤s j≤k) = 57 | H , n , ⇓ᵏhalt-upClosed M⇓Rn-k H (∸-monoʳ-≤ n (s≤s j≤k)) 58 | 59 | infix 8 _⇓ᵒ_ 60 | _⇓ᵒ_ : Term → Result → Setᵒ 61 | M ⇓ᵒ R = record { # = (M ⇓ R) 62 | ; down = downClosed⇓ M R 63 | ; tz = tt 64 | } 65 | 66 | _⇑ : Term → ℕ → Set 67 | (M ⇑) k = (M ⇓ᵏ timeout) k 68 | 69 | downClosed⇑ : ∀ M → downClosed (M ⇑) 70 | downClosed⇑ M k M⇑ j j≤k = ⇓ᵏtimeout-downClosed M⇑ j≤k 71 | 72 | infix 8 _⇑ᵒ 73 | _⇑ᵒ : Term → Setᵒ 74 | M ⇑ᵒ = record { # = (M ⇑) 75 | ; down = downClosed⇑ M 76 | ; tz = ⇓ᵏzero 77 | } 78 | 79 | infix 6 _⟹_ 80 | _⟹_ : Term → Result → ℕ → Set 81 | (M ⟹ val V) k = (M ⇓ val V) k 82 | (M ⟹ blameR) k = (M ⇓ blameR) k 83 | (M ⟹ timeout) k = (M ⇑) k 84 | 85 | downClosed⟹ : ∀ M R → downClosed (M ⟹ R) 86 | downClosed⟹ M (val V) = downClosed⇓ M (val V) 87 | downClosed⟹ M blameR = downClosed⇓ M blameR 88 | downClosed⟹ M timeout = downClosed⇑ M 89 | 90 | trueZero⟹ : ∀ M R → (M ⟹ R) 0 91 | trueZero⟹ M (val V) = tt 92 | trueZero⟹ M blameR = tt 93 | trueZero⟹ M timeout = ⇓ᵏzero 94 | 95 | infix 8 _⟹ᵒ_ 96 | _⟹ᵒ_ : Term → Result → Setᵒ 97 | M ⟹ᵒ R = record { # = (M ⟹ R) 98 | ; down = downClosed⟹ M R 99 | ; tz = trueZero⟹ M R 100 | } 101 | 102 | ⟹E : ∀ M R {k} {P : Set} 103 | → (M ⟹ R) k 104 | → ((M ⇓ R) k → P) 105 | → ((M ⇑) k → P) 106 | → P 107 | ⟹E M (val V) M⟹R cont⇓ cont⇑ = cont⇓ M⟹R 108 | ⟹E M blameR M⟹R cont⇓ cont⇑ = cont⇓ M⟹R 109 | ⟹E M timeout M⟹R cont⇓ cont⇑ = cont⇑ M⟹R 110 | 111 | {- 112 | values-dont-diverge : ∀{V}{k} 113 | → Value V 114 | → (V ⇑) (suc k) 115 | → ⊥ 116 | values-dont-diverge{V}{k} (v 〈 _ 〉) (inj⇓ᵏ-raise V⇑ x) = 117 | values-dont-diverge{k = k} v {!!} 118 | -} 119 | -------------------------------------------------------------------------------- /src/rewriting/examples/ClosureConversion.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --rewriting #-} 2 | {- 3 | This is a language without lexical scoping, but otherwise similar to the lambda calculus. 4 | -} 5 | 6 | open import Data.List using (List; []; _∷_; length) 7 | open import Data.Nat using (ℕ; zero; suc; _≟_) 8 | open import Relation.Nullary using (Dec; yes; no) 9 | open import Data.List.Membership.Propositional 10 | open import Data.List.Relation.Unary.Any 11 | open import Data.Product using (_×_; _,_; proj₁; proj₂; ∃-syntax) 12 | open import Data.Unit.Polymorphic using (⊤; tt) 13 | open import Data.Vec using (Vec) renaming ([] to []̌; _∷_ to _∷̌_) 14 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym) 15 | open import Sig 16 | open import Var 17 | open import rewriting.examples.Denot 18 | open import rewriting.examples.Gamma 19 | renaming (Term to Term₁; $ to $₁; _·_ to _·₁_; 20 | ⟨_,_⟩ to ⟨_,_⟩₁; fst to fst₁; snd to snd₁; 21 | ⟦_⟧ to ⟦_⟧₁) 22 | open import rewriting.examples.Delta 23 | using (δ) 24 | renaming (Term to Term₂; $ to $₂; _·_ to _·₂_; 25 | ⟨_,_⟩ to ⟨_,_⟩₂; fst to fst₂; snd to snd₂; 26 | `_ to `₂_; ⟦_⟧ to ⟦_⟧₂) 27 | {- 28 | open import Data.List.Membership.DecPropositional _≟_ 29 | -} 30 | 31 | module rewriting.examples.ClosureConversion where 32 | 33 | {- Closure Conversion -} 34 | C : Term₁ → Term₂ 35 | C (` x) = `₂ x 36 | C ($₁ k) = $₂ k 37 | C (γ N M) = ⟨ δ (C N) , C M ⟩₂ 38 | C (L ·₁ M) = ((fst₂ (C L)) ·₂ snd₂ (C L)) ·₂ (C M) 39 | C ⟨ L , M ⟩₁ = ⟨ C L , C M ⟩₂ 40 | C (fst₁ L) = fst₂ (C L) 41 | C (snd₁ L) = snd₂ (C L) 42 | 43 | continuous₂ : ∀ N {D ρ d} → ⟦ N ⟧₂ (D ∷ ρ) d → ∃[ V ] ⟦ N ⟧₂ ((_∈ V) ∷ ρ) d 44 | continuous₂ N {D}{ρ}{d} = {!!} 45 | 46 | data _≈_ : List Val → List Val → Set 47 | 48 | data _∼_ : Val → Val → Set where 49 | lit-rel : ∀ {k} 50 | → lit k ∼ lit k 51 | fun-rel : ∀ {V V′ V″ w w′ F} 52 | → w ∼ w′ 53 | → V ≈ V″ 54 | → V′ ↦ (V″ ↦ w′) ∈ F 55 | → (V ↦ w) ∼ pair F V′ 56 | 57 | data _≈_ where 58 | ≈-nil : [] ≈ [] 59 | ≈-cons : ∀ {v w}{vs ws} 60 | → v ∼ w 61 | → vs ≈ ws 62 | → (v ∷ vs) ≈ (w ∷ ws) 63 | 64 | C-fwd : ∀ M ρ d₁ → ⟦ M ⟧₁ ρ d₁ → ∃[ d₂ ] ⟦ C M ⟧₂ ρ d₂ × d₁ ∼ d₂ 65 | C-fwd (` x) ρ d₁ d₁∈M = {!!} , ({!!} , {!!}) 66 | C-fwd ($₁ k) ρ d₁ refl = (lit k) , refl , {!!} 67 | C-fwd (γ N M) ρ (V ↦ w) ((d , d∈M) , w∈⟦N⟧) 68 | with C-fwd M ρ d d∈M 69 | ... | d′ , d′∈CM , ∼d′ 70 | with C-fwd N (⟦ M ⟧₁ ρ ∷ (_∈ V) ∷ []) w w∈⟦N⟧ 71 | ... | d₂ , d₂∈CN , ∼d₂ 72 | with continuous₂ (C N) d₂∈CN 73 | ... | V′ , d₂∈CN′ = 74 | (pair ({!!} ↦ ({!!} ↦ d₂) ∷ []) {!!}) , 75 | (({!!} , {!!}) , fun-rel ∼d₂ {!!} (here refl)) 76 | C-fwd (L ·₁ M) ρ w ((v , v∈M) , (V , (Vw∈L , v∈V→v∈M))) 77 | with C-fwd L ρ (V ↦ w) Vw∈L 78 | ... | d′ , d′∈CL , fun-rel {F = F} w∼w′ V≈V″ ∈F = 79 | w′ , (neCM , (V′ , ((({!!} , ({!!} , ({!!} , (d′∈CL , {!!})))) , 80 | ({!!} , (({!!} , ({!!} , (d′∈CL , {!!}))) , {!!}))) , {!!}))) , 81 | w∼w′ 82 | {- 83 | w′ , ((neCM , V′ , ((neπ₂CL , V″ , VVw∈π₁CL , V″⊆π₂CL) , V′⊆CM)) , w∼w′) 84 | -} 85 | where 86 | neCM : ∃[ v ] ⟦ (C M) ⟧₂ ρ v 87 | neCM = {!!} 88 | 89 | V′ : List Val 90 | V′ = {!!} 91 | 92 | w′ : Val 93 | w′ = {!!} 94 | 95 | V′⊆CM : ∀ v₁ → v₁ ∈ V′ → ⟦ C M ⟧₂ ρ v₁ 96 | V′⊆CM = {!!} 97 | 98 | neπ₂CL : ∃[ v ] π₂ (⟦ C L ⟧₂ ρ) v 99 | neπ₂CL = _ , (_ , {!!}) 100 | 101 | V″ : List Val 102 | V″ = {!!} 103 | 104 | p₂ : Val 105 | p₂ = {!!} 106 | 107 | ∈CL : ⟦ C L ⟧₂ ρ (pair F _) 108 | ∈CL = d′∈CL 109 | 110 | VVw∈π₁CL : π₁ (⟦ C L ⟧₂ ρ) (V″ ↦ (V′ ↦ w′)) 111 | VVw∈π₁CL = {!!} , {!!} 112 | 113 | V″⊆π₂CL : ∀ v₁ → v₁ ∈ V″ → π₂ (⟦ C L ⟧₂ ρ) v₁ 114 | V″⊆π₂CL v₁ v₁∈V″ = {!!} , {!!} 115 | 116 | C-fwd ⟨ L , M ⟩₁ ρ d₁ d₁∈M = {!!} 117 | C-fwd (fst₁ L) ρ d₁ d₁∈M = {!!} 118 | C-fwd (snd₁ L) ρ d₁ d₁∈M = {!!} 119 | -------------------------------------------------------------------------------- /src/weaken/AbstractBindingTree.agda: -------------------------------------------------------------------------------- 1 | {- 2 | Based on Philip Wadler's paper Explicit Weakening. 3 | -} 4 | 5 | {-# OPTIONS --rewriting #-} 6 | open import Agda.Builtin.Equality 7 | open import Agda.Builtin.Equality.Rewrite 8 | 9 | open import Data.Nat using (ℕ; zero; suc; _+_; _⊔_; _∸_; _≟_) 10 | open import Data.List using (List; []; _∷_) 11 | import Relation.Binary.PropositionalEquality as Eq 12 | open Eq using (_≡_; _≢_; refl; sym; cong; cong₂; cong-app; subst) 13 | open Eq.≡-Reasoning 14 | 15 | module weaken.AbstractBindingTree {ℓ} (Op : Set ℓ) (sig : Op → List ℕ) where 16 | 17 | infixl 5 _▷_ 18 | infixl 5 _⨟_ 19 | infix 8 _↑ 20 | 21 | data Args : List ℕ → Set ℓ 22 | 23 | data ABT : Set ℓ where 24 | • : ABT 25 | _↑ : ABT → ABT 26 | _⦅_⦆ : (op : Op) → Args (sig op) → ABT 27 | 28 | data Arg : ℕ → Set ℓ where 29 | ast : ABT → Arg 0 30 | bind : ∀{b} → Arg b → Arg (suc b) 31 | 32 | data Args where 33 | nil : Args [] 34 | cons : ∀{b bs} → Arg b → Args bs → Args (b ∷ bs) 35 | 36 | data Subst : Set ℓ where 37 | id : Subst 38 | _↑ : Subst → Subst 39 | _▷_ : Subst → ABT → Subst 40 | 41 | _[_] : ABT → Subst → ABT 42 | _[_]ₐ : ∀{b} → Arg b → Subst → Arg b 43 | _[_]* : ∀{bs} → Args bs → Subst → Args bs 44 | M [ id ] = M 45 | M [ σ ↑ ] = M [ σ ] ↑ 46 | • [ σ ▷ N ] = N 47 | (M ↑) [ σ ▷ N ] = M [ σ ] 48 | (op ⦅ args ⦆) [ σ ▷ N ] = op ⦅ args [ σ ▷ N ]* ⦆ 49 | 50 | ast M [ σ ]ₐ = ast (M [ σ ]) 51 | bind arg [ σ ]ₐ = bind (arg [ σ ↑ ▷ • ]ₐ) 52 | 53 | nil [ σ ]* = nil 54 | (cons arg args) [ σ ]* = cons (arg [ σ ]ₐ) (args [ σ ]*) 55 | 56 | _⨟_ : Subst → Subst → Subst 57 | σ ⨟ id = σ 58 | σ ⨟ τ ↑ = (σ ⨟ τ) ↑ 59 | id ⨟ (τ ▷ N) = τ ▷ N 60 | σ ↑ ⨟ (τ ▷ N) = σ ⨟ τ 61 | (σ ▷ M) ⨟ (τ ▷ N) = σ ⨟ (τ ▷ N) ▷ (M [ τ ▷ N ]) 62 | 63 | sub-sub-compose : ∀(M : ABT) (σ τ : Subst) → (M [ σ ] [ τ ]) ≡ M [ σ ⨟ τ ] 64 | sub-sub-compose-arg : ∀{b}(arg : Arg b) (σ τ : Subst) → arg [ σ ]ₐ [ τ ]ₐ ≡ arg [ σ ⨟ τ ]ₐ 65 | sub-sub-compose-args : ∀{b}(args : Args b) (σ τ : Subst) → args [ σ ]* [ τ ]* ≡ args [ σ ⨟ τ ]* 66 | 67 | sub-sub-compose M σ id = refl 68 | sub-sub-compose M σ (τ ↑) = cong _↑ (sub-sub-compose M σ τ) 69 | sub-sub-compose M id (τ ▷ N) = refl 70 | sub-sub-compose M (σ ↑) (τ ▷ N) = sub-sub-compose M σ τ 71 | sub-sub-compose • (σ ▷ L) (τ ▷ N) = refl 72 | sub-sub-compose (M ↑) (σ ▷ L) (τ ▷ N) = sub-sub-compose M σ (τ ▷ N) 73 | sub-sub-compose (op ⦅ args ⦆) (σ ▷ L) (τ ▷ N) = 74 | cong (λ X → op ⦅ X ⦆) (sub-sub-compose-args args (σ ▷ L) (τ ▷ N)) 75 | 76 | sub-sub-compose-arg (ast M) σ τ = cong ast (sub-sub-compose M σ τ) 77 | sub-sub-compose-arg (bind arg) σ τ = cong bind (sub-sub-compose-arg arg (σ ↑ ▷ •) (τ ↑ ▷ •)) 78 | 79 | sub-sub-compose-args nil σ τ = refl 80 | sub-sub-compose-args (cons arg args) σ τ = 81 | cong₂ cons (sub-sub-compose-arg arg σ τ) (sub-sub-compose-args args σ τ) 82 | 83 | {-# REWRITE sub-sub-compose sub-sub-compose-arg sub-sub-compose-args #-} 84 | 85 | left-id : (τ : Subst) → id ⨟ τ ≡ τ 86 | left-id id = refl 87 | left-id (τ ↑) = cong _↑ (left-id τ) 88 | left-id (τ ▷ N) = refl 89 | 90 | {-# REWRITE left-id #-} 91 | 92 | assoc : (σ τ ν : Subst) → (σ ⨟ τ) ⨟ ν ≡ σ ⨟ (τ ⨟ ν) 93 | assoc σ τ id = refl 94 | assoc σ τ (ν ↑) = cong _↑ (assoc σ τ ν) 95 | assoc σ id (ν ▷ N) = refl 96 | assoc σ (τ ↑) (ν ▷ N) = assoc σ τ ν 97 | assoc id (τ ▷ M) (ν ▷ N) = refl 98 | assoc (σ ↑) (τ ▷ M) (ν ▷ N) = assoc σ τ (ν ▷ N) 99 | assoc (σ ▷ L) (τ ▷ M) (ν ▷ N) = cong₂ _▷_ (assoc σ (τ ▷ M) (ν ▷ N)) refl 100 | 101 | {-# REWRITE assoc #-} 102 | 103 | _[_]₀ : ABT → ABT → ABT 104 | N [ M ]₀ = N [ id ▷ M ] 105 | 106 | _[_]₁ : ABT → ABT → ABT 107 | N [ M ]₁ = N [ (id ▷ M) ↑ ▷ • ] 108 | 109 | _[_]₁[_]₀ : ABT → ABT → ABT → ABT 110 | N [ M ]₁[ L ]₀ = N [ id ▷ M ▷ L ] 111 | 112 | up-subst-id : ∀(M N : ABT) → (N ↑) [ M ]₀ ≡ N 113 | up-subst-id M N = refl 114 | 115 | commute-subst : ∀(L M N : ABT) → N [ M ]₀ [ L ]₀ ≡ N [ L ]₁ [ M [ L ]₀ ]₀ 116 | commute-subst L M N = refl 117 | 118 | exts-subst-cons : ∀(σ : Subst)(N V : ABT) → N [ σ ↑ ▷ • ] [ V ]₀ ≡ N [ σ ▷ V ] 119 | exts-subst-cons σ N V = refl 120 | -------------------------------------------------------------------------------- /src/rewriting/examples/SystemF.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | 3 | open import Data.Empty using (⊥; ⊥-elim) 4 | open import Data.List using (List; []; _∷_; length; map; _++_) 5 | open import Data.Nat using (ℕ; zero; suc; _+_; _≤_; _≤?_) 6 | open import Data.Product using (_×_; proj₁; proj₂; ∃-syntax) 7 | renaming (_,_ to ⟨_,_⟩ ) 8 | open import Data.Unit.Polymorphic using (⊤; tt) 9 | open import Data.Vec using (Vec) renaming ([] to []̌; _∷_ to _∷̌_) 10 | open import Relation.Binary.PropositionalEquality as Eq 11 | using (_≡_; refl; sym; trans; cong; cong₂; subst) 12 | open import Function using (_∘_) 13 | open import Sig 14 | 15 | module rewriting.examples.SystemF where 16 | 17 | {------------- Types -------------} 18 | 19 | 20 | data TypeOp : Set where 21 | op-fun : TypeOp 22 | op-all : TypeOp 23 | op-nat-ty : TypeOp 24 | 25 | type-sig : TypeOp → List Sig 26 | type-sig op-fun = ■ ∷ ■ ∷ [] 27 | type-sig op-all = (ν ■) ∷ [] 28 | type-sig op-nat-ty = [] 29 | 30 | open import rewriting.AbstractBindingTree TypeOp type-sig 31 | using () 32 | renaming (ABT to Type; Rename to Renameᵗ; Subst to Substᵗ; 33 | ren to renᵗ; ren-def to ren-defᵗ; extr to extrᵗ; ext to extᵗ; 34 | ⟪_⟫ to ⟪_⟫ᵗ; sub-var to sub-varᵗ; seq-def to seq-defᵗ; ↑ to ↑ᵗ; 35 | _[_] to _⦗_⦘; _⦅_⦆ to _‹_›; _•_ to _•ᵗ_; id to idᵗ; _⨟_ to _⨟ᵗ_; 36 | nil to tnil; cons to tcons; bind to tbind; ast to tast; `_ to ^_) public 37 | 38 | pattern Nat = op-nat-ty ‹ tnil › 39 | 40 | infixl 7 _⇒_ 41 | pattern _⇒_ A B = op-fun ‹ tcons (tast A) (tcons (tast B) tnil) › 42 | 43 | pattern ∀̇ A = op-all ‹ tcons (tbind (tast A)) tnil › 44 | 45 | {------------- Terms -------------} 46 | 47 | data Op : Set where 48 | op-nat : ℕ → Op 49 | op-lam : Op 50 | op-app : Op 51 | op-tyabs : Op 52 | op-tyapp : Op 53 | 54 | sig : Op → List Sig 55 | sig (op-nat n) = [] 56 | sig op-lam = (ν ■) ∷ [] 57 | sig op-app = ■ ∷ ■ ∷ [] 58 | sig op-tyabs = ■ ∷ [] 59 | sig op-tyapp = ■ ∷ [] 60 | 61 | open import rewriting.AbstractBindingTree Op sig renaming (ABT to Term) public 62 | 63 | pattern $ n = (op-nat n) ⦅ nil ⦆ 64 | pattern ƛ N = op-lam ⦅ cons (bind (ast N)) nil ⦆ 65 | infixl 7 _·_ 66 | pattern _·_ L M = op-app ⦅ cons (ast L) (cons (ast M) nil) ⦆ 67 | pattern Λ N = op-tyabs ⦅ cons (ast N) nil ⦆ 68 | pattern _[·] L = op-tyapp ⦅ cons (ast L) nil ⦆ 69 | 70 | {----------------------- Values ------------------------} 71 | 72 | data Value : Term → Set where 73 | 74 | V-nat : ∀ {n : ℕ} 75 | ------------- 76 | → Value ($ n) 77 | 78 | V-ƛ : ∀ {N : Term} 79 | --------------------------- 80 | → Value (ƛ N) 81 | 82 | V-Λ : ∀ {N : Term} 83 | --------------------------- 84 | → Value (Λ N) 85 | 86 | value : ∀ {V : Term} 87 | → (v : Value V) 88 | ------------- 89 | → Term 90 | value {V = V} v = V 91 | 92 | {----------------------- Frames ------------------------} 93 | 94 | infix 6 □·_ 95 | infix 6 _·□ 96 | 97 | data Frame : Set where 98 | 99 | □·_ : 100 | (M : Term) 101 | ---------- 102 | → Frame 103 | 104 | _·□ : ∀ {V : Term} 105 | → (v : Value V) 106 | ------------- 107 | → Frame 108 | 109 | □[·] : Frame 110 | 111 | ƛ□ : Frame 112 | 113 | _⟦_⟧ : Frame → Term → Term 114 | (□· M) ⟦ L ⟧ = L · M 115 | (v ·□) ⟦ M ⟧ = value v · M 116 | □[·] ⟦ M ⟧ = M [·] 117 | ƛ□ ⟦ M ⟧ = ƛ M 118 | 119 | {------------- Reduction Semantics -------------} 120 | 121 | infix 2 _—→_ 122 | 123 | data _—→_ : Term → Term → Set where 124 | 125 | ξξ : ∀ {M N : Term} {M′ N′ : Term} 126 | → (F : Frame) 127 | → M′ ≡ F ⟦ M ⟧ 128 | → N′ ≡ F ⟦ N ⟧ 129 | → M —→ N 130 | -------- 131 | → M′ —→ N′ 132 | 133 | β-ƛ : ∀ {N M : Term} 134 | -------------------- 135 | → (ƛ N) · M —→ N [ M ] 136 | 137 | β-Λ : ∀ {N : Term} 138 | ------------------ 139 | → (Λ N) [·] —→ N 140 | 141 | pattern ξ F M—→N = ξξ F refl refl M—→N 142 | 143 | -------------------------------------------------------------------------------- /isabelle/junk/LambdaExample.thy: -------------------------------------------------------------------------------- 1 | theory LambdaExample 2 | imports Main AbstractBindingTrees 3 | begin 4 | 5 | section "Lambda Calculus" 6 | 7 | datatype op_lam = LamOp | AppOp 8 | 9 | type_synonym Term = "op_lam ABT" 10 | 11 | abbreviation Lam :: "Term \ Term" ("\") where 12 | "\ N \ App LamOp [Bnd (Trm N)]" 13 | 14 | abbreviation Apply :: "Term \ Term \ Term" (infixl "\" 60) where 15 | "L \ M \ App AppOp [Trm L, Trm M]" 16 | 17 | inductive WF_op :: "op_lam \ unit list list \ unit list \ unit \ bool" where 18 | WFLamOp[intro!]: "WF_op LamOp [[()]] [()] ()" | 19 | WFAppOp[intro!]: "WF_op AppOp [[],[]] [(),()] ()" 20 | 21 | context abt_predicate 22 | begin 23 | 24 | inductive_cases 25 | wf_var_elim[elim!]: "\ \ Var x : T" and 26 | wf_app_elim[elim!]: "\ \ App op args : T" and 27 | wf_trm_elim[elim!]: "\; Bs \\<^sub>a Trm M : T" and 28 | wf_bnd_elim[elim!]: "\; Bs \\<^sub>a Bnd A : T" and 29 | wf_cons_elim[elim!]: "\; Bss \\<^sub>+ As : Ts" 30 | 31 | end 32 | 33 | interpretation lam: abt_predicate WF_op . 34 | 35 | abbreviation WF :: "unit list \ Term \ bool" ("_\_") where 36 | "\ \ M \ lam.wf_abt \ M ()" 37 | 38 | inductive wf :: "unit list \ Term \ bool" where 39 | wf_var[intro!]: "nth \ x = () \ wf \ (Var x)" | 40 | wf_app[intro!]: "\ wf \ L; wf \ M \ \ wf \ (L \ M)" | 41 | wf_abs[intro!]: "wf (()#\) N \ wf \ (\ N)" 42 | 43 | lemma wf_implies_WF: "wf \ M \ \ \ M" 44 | by (induction \ M rule: wf.induct) auto 45 | 46 | thm ABT.induct 47 | thm wf.induct 48 | thm lam.wf_abt_wf_args_wf_arg.induct 49 | 50 | lemma WF_implies_wf: "\ \ M \ wf \ M" 51 | apply (induction \ M rule: lam.wf_abt_wf_args_wf_arg.induct) 52 | sorry 53 | 54 | inductive reduce :: "Term \ Term \ bool" (infix "\" 50) where 55 | beta[intro!]: "(\ N) \ M \ N[M]" | 56 | appl[intro!]: "L \ L' \ L \ M \ L' \ M" | 57 | appr[intro!]: "M \ M' \ L \ M \ L \ M'" | 58 | abs[intro!]: "N \ N' \ \ N \ \ N'" 59 | 60 | lemma reduce_WF: "\ M \ N; \ \ M \ \ \ \ N" 61 | proof (induction arbitrary: \ rule: reduce.induct) 62 | case (beta N M) 63 | have 1: "\ \ M" using beta by auto 64 | have 2: "()#\ \ N" using beta by auto 65 | from 1 2 show ?case apply auto apply (rule lam.subst_pres) apply force 66 | unfolding lam.sub_pres_def \_def apply auto apply (case_tac x) apply auto done 67 | next 68 | case (appl L L' M) 69 | then show ?case by auto 70 | next 71 | case (appr M M' L) 72 | then show ?case by auto 73 | next 74 | case (abs N N' \) 75 | have "()#\ \ N" using abs by auto 76 | then have "()#\ \ N'" using abs by blast 77 | then show ?case by auto 78 | qed 79 | 80 | 81 | section "Confluence" 82 | 83 | inductive reduce_parallel :: "Term \ Term \ bool" (infix "\" 50) where 84 | pvar[intro!]: "Var x \ Var x" | 85 | pabs[intro!]: "N \ N' \ \ N \ \ N'" | 86 | papp[intro!]: "\ L \ L'; M \ M' \ \ L \ M \ L' \ M'" | 87 | pbeta[intro!]: "\ N \ N'; M \ M' \ \ (\ N) \ M \ N'[M']" 88 | 89 | (* 90 | lemma parallel_refl: "\ \ M \ M \ M" 91 | by (induction rule: WF.induct) auto 92 | *) 93 | 94 | lemma beta_parallel: "M \ N \ M \ N" 95 | apply (induction rule: reduce.induct) 96 | sorry 97 | 98 | end -------------------------------------------------------------------------------- /src/experimental/MapPreserve.agda: -------------------------------------------------------------------------------- 1 | {--------------------------- 2 | 3 | Preservation of a predicate 4 | 5 | Let "I" be the kind for type-like things. 6 | 7 | A : I 8 | Γ Δ : List I 9 | 10 | preserve-map : ∀{M σ Γ Δ A} 11 | → Γ ⊢ M ⦂ A 12 | → σ ⦂ Γ ⇒ Δ 13 | → Δ ⊢ map-abt σ M ⦂ A 14 | 15 | ---------------------------} 16 | 17 | import ABTPredicate 18 | open import Agda.Primitive using (Level; lzero; lsuc) 19 | open import Data.Empty using (⊥) 20 | open import Data.List using (List; []; _∷_; length; _++_) 21 | open import Data.Nat using (ℕ; zero; suc; _+_; _<_; z≤n; s≤s) 22 | open import Data.Nat.Properties using (≤-refl) 23 | open import Data.Product using (_×_; proj₁; proj₂; Σ-syntax) 24 | renaming (_,_ to ⟨_,_⟩ ) 25 | open import Data.Unit.Polymorphic using (⊤; tt) 26 | open import Function using (_∘_) 27 | import Substitution 28 | open import Structures 29 | open import GSubst 30 | open import GenericSubstitution 31 | open import Data.Vec using (Vec) renaming ([] to []̌; _∷_ to _∷̌_) 32 | import Relation.Binary.PropositionalEquality as Eq 33 | open Eq using (_≡_; refl; sym; trans; cong; cong₂; cong-app) 34 | open Eq.≡-Reasoning 35 | open import Var 36 | 37 | module MapPreserve (Op : Set) (sig : Op → List ℕ) where 38 | 39 | open import AbstractBindingTree Op sig 40 | open import Map Op sig 41 | 42 | record MapPreservable (V I : Set){{_ : Quotable V}} {{_ : Shiftable V}} : Set₁ 43 | where 44 | field 𝑉 : List I → Var → I → Set 45 | 𝑃 : (op : Op) → Vec I (length (sig op)) → BTypes I (sig op) → I → Set 46 | _⊢v_⦂_ : List I → V → I → Set 47 | ⊢v0 : ∀{B Δ} → (B ∷ Δ) ⊢v var→val 0 ⦂ B 48 | shift-⊢v : ∀{A B Δ v} → Δ ⊢v v ⦂ A → (B ∷ Δ) ⊢v ⇑ v ⦂ A 49 | open ABTPredicate Op sig 𝑉 𝑃 public 50 | field quote-⊢v : ∀{Γ v A} → Γ ⊢v v ⦂ A → Γ ⊢ “ v ” ⦂ A 51 | 52 | open MapPreservable {{...}} 53 | 54 | _⦂_⇒_ : ∀{V I : Set} 55 | {{_ : Quotable V}} {{_ : Shiftable V}} {{_ : MapPreservable V I}} 56 | → GSubst V → List I → List I → Set 57 | _⦂_⇒_ {V}{I} σ Γ Δ = ∀{x : Var} {A : I} → Γ ∋ x ⦂ A → Δ ⊢v σ x ⦂ A 58 | 59 | 60 | preserve-map : ∀ {V I : Set}{Γ Δ : List I}{σ : GSubst V}{A : I} 61 | {{_ : Quotable V}} {{_ : Shiftable V}} {{_ : MapPreservable V I}} 62 | {{_ : Renameable V}} 63 | (M : ABT) 64 | → Γ ⊢ M ⦂ A 65 | → σ ⦂ Γ ⇒ Δ 66 | → Δ ⊢ map σ M ⦂ A 67 | preserve-map (` x) (var-p ∋x 𝑉x) σ⦂ = quote-⊢v (σ⦂ ∋x) 68 | preserve-map {V}{I} (op ⦅ args ⦆) (op-p ⊢args Pop) σ⦂ = 69 | op-p (pres-args ⊢args σ⦂) Pop 70 | where 71 | map-pres-ext : ∀ {σ : GSubst V}{Γ Δ : List I}{A : I} 72 | → σ ⦂ Γ ⇒ Δ → ext σ ⦂ (A ∷ Γ) ⇒ (A ∷ Δ) 73 | map-pres-ext {σ = σ} σ⦂ {zero} refl = ⊢v0 74 | map-pres-ext {σ = σ} σ⦂ {suc x} ∋x = shift-⊢v (σ⦂ ∋x) 75 | 76 | pres-arg : ∀{b Γ Δ}{arg : Arg b}{A σ Bs} 77 | → b ∣ Γ ∣ Bs ⊢ₐ arg ⦂ A → σ ⦂ Γ ⇒ Δ 78 | → b ∣ Δ ∣ Bs ⊢ₐ map-arg σ {b} arg ⦂ A 79 | pres-args : ∀{bs Γ Δ}{args : Args bs}{As σ Bss} 80 | → bs ∣ Γ ∣ Bss ⊢₊ args ⦂ As → σ ⦂ Γ ⇒ Δ 81 | → bs ∣ Δ ∣ Bss ⊢₊ map-args σ {bs} args ⦂ As 82 | pres-arg {b} {arg = ast M} (ast-p ⊢M) σ⦂ = 83 | ast-p (preserve-map M ⊢M σ⦂) 84 | pres-arg {suc b}{Γ}{Δ}{bind arg}{σ = σ} (bind-p {B = B}{A = A} ⊢arg) σ⦂ = 85 | bind-p (pres-arg ⊢arg (λ{x}{A} → map-pres-ext{σ}{Γ}{Δ} σ⦂ {x}{A})) 86 | pres-arg {b}{Γ₁}{Δ₁}{arg = perm f f⁻¹ inv inv' arg}{A}{σ₁}{Bs} 87 | (perm-p{Γ = Γ₁}{Δ = Γ₂} _ _ fbnd Δ= ⊢arg) σ⦂ = 88 | let Δ₂ = ren-ctx f⁻¹ Δ₁ (length Δ₁) ≤-refl {!!} in 89 | let IH = pres-arg{σ = ren f ∘ σ₁ ∘ f⁻¹} ⊢arg {!!} in 90 | perm-p{Γ = Δ₁}{Δ = {!!}} {!!} {!!} {!!} {!!} IH 91 | {- Have: 92 | σ₁ ⦂ Γ₁ ⇒ Δ₁ 93 | Γ₂ = ren-ctx f⁻¹ Γ₁ 94 | b ∣ Γ₂ ∣ Bs₁ ⊢ₐ arg ⦂ A 95 | 96 | Γ₁ -- σ₁ --> Δ₁ 97 | |^ |^ 98 | f| f| 99 | |f⁻¹ |f⁻¹ 100 | V| V| 101 | Γ₂ .. σ₂ ..> Δ₂ 102 | 103 | Need: 104 | σ₂ ⦂ Γ₂ ⇒ Δ₂ 105 | σ₂ = ren f ∘ σ₁ ∘ f⁻¹ 106 | 107 | -} 108 | 109 | 110 | pres-args {[]} {args = nil} nil-p σ⦂ = nil-p 111 | pres-args {b ∷ bs} {args = cons arg args} (cons-p ⊢arg ⊢args) σ⦂ = 112 | cons-p (pres-arg ⊢arg σ⦂) (pres-args ⊢args σ⦂) 113 | -------------------------------------------------------------------------------- /src/rewriting/examples/CastDiverge.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module rewriting.examples.CastDiverge where 3 | 4 | open import Data.List using (List; []; _∷_; length; map) 5 | open import Data.Maybe 6 | open import Data.Nat 7 | open import Data.Bool using (true; false) renaming (Bool to 𝔹) 8 | open import Data.Nat.Properties 9 | open import Data.Product using (_,_;_×_; proj₁; proj₂; Σ-syntax; ∃-syntax) 10 | open import Data.Unit using (⊤; tt) 11 | open import Data.Unit.Polymorphic renaming (⊤ to topᵖ; tt to ttᵖ) 12 | open import Data.Empty using (⊥; ⊥-elim) 13 | open import Data.Sum using (_⊎_; inj₁; inj₂) 14 | open import Relation.Binary.PropositionalEquality as Eq 15 | using (_≡_; _≢_; refl; sym; cong; subst; trans) 16 | open import Relation.Nullary using (¬_; Dec; yes; no) 17 | open import Var 18 | open import rewriting.examples.Cast 19 | open import rewriting.examples.CastBigStep 20 | open import rewriting.examples.StepIndexedLogic2 21 | 22 | {----- Diverge ------} 23 | 24 | data ⇑ : Term → ℕ → Set where 25 | ⇑zero : ∀{M} → ⇑ M zero 26 | ⇑app : ∀{L M N W k} 27 | → L ⇓ ƛ N 28 | → M ⇓ W → Value W 29 | → ⇑ (N [ W ]) k 30 | → ⇑ (L · M) (suc k) 31 | ⇑app-L : ∀{L M k} → ⇑ L k → ⇑ (L · M) (suc k) 32 | ⇑app-R : ∀{L M N k} → L ⇓ ƛ N → ⇑ M k → ⇑ (L · M) (suc k) 33 | ⇑inj : ∀{M G k} → ⇑ M k → ⇑ (M ⟨ G !⟩) (suc k) 34 | ⇑proj : ∀{M H k} → ⇑ M k → ⇑ (M ⟨ H ?⟩) (suc k) 35 | 36 | downClosed⇑ : ∀ {M} → downClosed (⇑ M) 37 | downClosed⇑ zero ⇑M .zero z≤n = ⇑zero 38 | downClosed⇑ (suc k) ⇑M .zero z≤n = ⇑zero 39 | downClosed⇑ (suc k) (⇑app L⇓ M⇓ w ⇑NW) (suc j) (s≤s j≤k) = 40 | ⇑app L⇓ M⇓ w (downClosed⇑ k ⇑NW j j≤k) 41 | downClosed⇑ (suc k) (⇑app-L ⇑M) (suc j) (s≤s j≤k) = 42 | ⇑app-L (downClosed⇑ k ⇑M j j≤k) 43 | downClosed⇑ (suc k) (⇑app-R x ⇑M) (suc j) (s≤s j≤k) = 44 | ⇑app-R x (downClosed⇑ k ⇑M j j≤k) 45 | downClosed⇑ (suc k) (⇑inj ⇑M) (suc j) (s≤s j≤k) = 46 | ⇑inj (downClosed⇑ k ⇑M j j≤k) 47 | downClosed⇑ (suc k) (⇑proj ⇑M) (suc j) (s≤s j≤k) = 48 | ⇑proj (downClosed⇑ k ⇑M j j≤k) 49 | 50 | {----- Diverge in SIL ------} 51 | 52 | ⇑ᵒ : Term → Setᵒ 53 | ⇑ᵒ M = record { # = ⇑ M 54 | ; down = downClosed⇑ 55 | ; tz = ⇑zero 56 | } 57 | 58 | {---- Lift Divergence Rules into SIL -----} 59 | 60 | ⊢ᵒ⇑app-L : ∀{𝒫}{L}{M} 61 | → 𝒫 ⊢ᵒ ▷ᵒ (⇑ᵒ L) 62 | → 𝒫 ⊢ᵒ ⇑ᵒ (L · M) 63 | ⊢ᵒ⇑app-L {𝒫}{L}{M} ⊢▷⇑L = ⊢ᵒ-intro 64 | λ { zero 𝒫z → ⇑zero 65 | ; (suc n) 𝒫sn → ⇑app-L (⊢ᵒ-elim ⊢▷⇑L (suc n) 𝒫sn) } 66 | 67 | ⊢ᵒ⇑app-R : ∀{𝒫}{L}{M}{N} 68 | → 𝒫 ⊢ᵒ (L ⇓ ƛ N)ᵒ 69 | → 𝒫 ⊢ᵒ ▷ᵒ (⇑ᵒ M) 70 | → 𝒫 ⊢ᵒ ⇑ᵒ (L · M) 71 | ⊢ᵒ⇑app-R {𝒫}{L}{M}{N} ⊢L⇓ ⊢▷⇑M = ⊢ᵒ-intro 72 | λ { zero _ → ⇑zero 73 | ; (suc n) 𝒫sn → 74 | ⇑app-R (⊢ᵒ-elim ⊢L⇓ (suc n) 𝒫sn) (⊢ᵒ-elim ⊢▷⇑M (suc n) 𝒫sn ) } 75 | 76 | ⊢ᵒ⇑app : ∀{𝒫}{L}{M}{N}{W} 77 | → 𝒫 ⊢ᵒ (L ⇓ ƛ N)ᵒ 78 | → 𝒫 ⊢ᵒ (M ⇓ W)ᵒ 79 | → 𝒫 ⊢ᵒ (Value W)ᵒ 80 | → 𝒫 ⊢ᵒ ▷ᵒ (⇑ᵒ (N [ W ])) 81 | → 𝒫 ⊢ᵒ ⇑ᵒ (L · M) 82 | ⊢ᵒ⇑app {𝒫}{L}{M}{N} ⊢L⇓ ⊢M⇓ ⊢w ⊢▷⇑NW = ⊢ᵒ-intro 83 | λ { zero _ → ⇑zero 84 | ; (suc n) 𝒫sn → 85 | ⇑app (⊢ᵒ-elim ⊢L⇓ (suc n) 𝒫sn) 86 | (⊢ᵒ-elim ⊢M⇓ (suc n) 𝒫sn) 87 | (⊢ᵒ-elim ⊢w (suc n) 𝒫sn) 88 | (⊢ᵒ-elim ⊢▷⇑NW (suc n) 𝒫sn) } 89 | 90 | ⊢⇑app-inv : ∀{𝒫}{L}{M}{R} 91 | → (▷ᵒ (⇑ᵒ L) ∷ 𝒫 ⊢ᵒ R) 92 | → (∀ N → (L ⇓ ƛ N)ᵒ ∷ ▷ᵒ (⇑ᵒ M) ∷ 𝒫 ⊢ᵒ R) 93 | → (∀ N W → (L ⇓ ƛ N)ᵒ ∷ (M ⇓ W)ᵒ ∷ (Value W)ᵒ ∷ ▷ᵒ (⇑ᵒ (N [ W ])) 94 | ∷ 𝒫 ⊢ᵒ R) 95 | → ⇑ᵒ (L · M) ∷ 𝒫 ⊢ᵒ R 96 | ⊢⇑app-inv {𝒫}{L}{M}{R} c1 c2 c3 = 97 | ⊢ᵒ-intro λ 98 | { zero x → tz R 99 | ; (suc n) (⇑app-L ⇑Ln , 𝒫sn) → 100 | ⊢ᵒ-elim c1 (suc n) (⇑Ln , 𝒫sn) 101 | ; (suc n) (⇑app-R L⇓ ⇑Mn , 𝒫sn) → 102 | ⊢ᵒ-elim (c2 _) (suc n) (L⇓ , ⇑Mn , 𝒫sn) 103 | ; (suc n) (⇑app L⇓ M⇓ w ⇑NWn , 𝒫sn) → 104 | ⊢ᵒ-elim (c3 _ _) (suc n) (L⇓ , M⇓ , w , ⇑NWn , 𝒫sn) 105 | } 106 | 107 | ⊢⇑inj : ∀{𝒫}{M}{G} 108 | → 𝒫 ⊢ᵒ ▷ᵒ (⇑ᵒ M) 109 | → 𝒫 ⊢ᵒ ⇑ᵒ (M ⟨ G !⟩) 110 | ⊢⇑inj {𝒫}{M}{G} ⊢▷⇑M = ⊢ᵒ-intro 111 | λ { zero 𝒫n → ⇑zero 112 | ; (suc n) 𝒫n → ⇑inj (⊢ᵒ-elim ⊢▷⇑M (suc n) 𝒫n)} 113 | 114 | ⊢⇑inj-inv : ∀{𝒫}{M}{G}{R} 115 | → ▷ᵒ (⇑ᵒ M) ∷ 𝒫 ⊢ᵒ R 116 | → ⇑ᵒ (M ⟨ G !⟩) ∷ 𝒫 ⊢ᵒ R 117 | ⊢⇑inj-inv {𝒫}{M}{G}{R} ▷⇑M⊢R = ⊢ᵒ-intro 118 | λ { zero _ → tz R 119 | ; (suc n) (⇑inj ⇑Mn , 𝒫sn) → ⊢ᵒ-elim ▷⇑M⊢R (suc n) (⇑Mn , 𝒫sn) } 120 | -------------------------------------------------------------------------------- /src/GSubst.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | open import Data.Nat using (ℕ; zero; suc; _+_) 3 | open import Data.Nat.Properties using (+-comm; +-assoc) 4 | open import Structures 5 | import Relation.Binary.PropositionalEquality as Eq 6 | open Eq using (_≡_; refl; sym; cong; cong₂; cong-app) 7 | open import Var 8 | 9 | module GSubst where 10 | 11 | GSubst : ∀{ℓ} (V : Set ℓ) → Set ℓ 12 | GSubst V = Var → V 13 | 14 | ↑ : (k : ℕ) → ∀{ℓ}{V : Set ℓ}{{_ : Shiftable V}} → GSubst {ℓ} V 15 | ↑ k x = var→val (k + x) 16 | 17 | infixr 6 _•_ 18 | _•_ : ∀{ℓ}{V : Set ℓ} → V → GSubst V → GSubst V 19 | (v • σ) 0 = v 20 | (v • σ) (suc x) = σ x 21 | 22 | ⟰ : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} → GSubst V → GSubst V 23 | ⟰ σ x = ⇑ (σ x) 24 | 25 | id : ∀ {ℓ}{V : Set ℓ}{{_ : Shiftable V}} → GSubst V 26 | id = ↑ 0 27 | 28 | ext : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} → GSubst V → GSubst V 29 | ext σ = (var→val 0) • ⟰ σ 30 | 31 | {- obsolete, use (v • ⟰ σ) instead of (σ , v) -} 32 | {- 33 | _,_ : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} → GSubst V → V → GSubst V 34 | (σ , v) = v • ⟰ σ 35 | -} 36 | 37 | drop : ∀{ℓ}{V : Set ℓ} → (k : ℕ) → GSubst V → GSubst V 38 | drop k σ x = σ (k + x) 39 | 40 | map-sub : ∀{ℓ}{V W : Set ℓ} → (V → W) → GSubst V → GSubst W 41 | map-sub f σ x = f (σ x) 42 | 43 | map-sub-id : ∀{ℓ}{V : Set ℓ} (σ : GSubst V) → map-sub (λ x → x) σ ≡ σ 44 | map-sub-id σ = refl 45 | 46 | map-sub-drop : ∀{ℓ} {V W : Set ℓ} σ f k 47 | → map-sub {ℓ}{V}{W} f (drop k σ) ≡ drop k (map-sub f σ) 48 | map-sub-drop σ f k = refl 49 | 50 | drop-0 : ∀ {ℓ}{V : Set ℓ} σ → drop {ℓ}{V} 0 σ ≡ σ 51 | drop-0 σ = refl 52 | 53 | drop-drop : ∀ {ℓ}{V} k k' σ → drop {ℓ} {V} (k + k') σ ≡ drop k (drop k' σ) 54 | drop-drop k k' σ = extensionality G 55 | where 56 | G : (x : Var) → drop (k + k') σ x ≡ drop k (drop k' σ) x 57 | G x rewrite +-comm k k' | +-assoc k' k x = refl 58 | 59 | shifts : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} → ℕ → V → V 60 | shifts zero v = v 61 | shifts (suc k) v = ⇑ (shifts k v) 62 | 63 | drop-inc : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} 64 | → (k : ℕ) (σ : GSubst V) → drop k (⟰ σ) ≡ ⟰ (drop k σ) 65 | drop-inc k σ = refl 66 | 67 | Z-shift : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} 68 | → (x : Var) → (var→val{ℓ}{V} 0 • ↑ 1) x ≡ var→val{ℓ}{V} x 69 | Z-shift 0 = refl 70 | Z-shift (suc x) = refl 71 | 72 | ext-cong : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} {σ₁ σ₂ : GSubst V} 73 | → ((x : ℕ) → σ₁ x ≡ σ₂ x) 74 | → (x : ℕ) → ext σ₁ x ≡ ext σ₂ x 75 | ext-cong {σ₁ = σ₁} {σ₂} f zero = refl 76 | ext-cong {σ₁ = σ₁} {σ₂} f (suc x) rewrite f x = refl 77 | 78 | drop-ext : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} 79 | → (k : Var) (σ : GSubst V) 80 | → drop (suc k) (ext σ) ≡ ⟰ (drop k σ) 81 | drop-ext k σ = refl 82 | 83 | Shift : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} → ℕ → GSubst V → Set ℓ 84 | Shift k σ = ∀ x → σ x ≡ shifts k (var→val x) 85 | 86 | inc-Shift : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} {k}{σ : GSubst V} 87 | → Shift k σ → Shift (suc k) (⟰ σ) 88 | inc-Shift {ℓ} {V} {k} {σ} shk x rewrite shk x = refl 89 | 90 | shifts-var→val : ∀{ℓ}{V : Set ℓ} {{S : Shiftable V}} 91 | → (k x : ℕ) → shifts k (var→val{ℓ}{V} x) ≡ var→val (k + x) 92 | shifts-var→val zero x = refl 93 | shifts-var→val{ℓ}{V}{{S}} (suc k) x rewrite shifts-var→val{ℓ}{V} k x 94 | | var→val-suc-shift {{S}} {x = k + x} = refl 95 | 96 | Shift-var : ∀{ℓ}{V : Set ℓ} {{_ : Shiftable V}} 97 | → (σ : GSubst V) (k : ℕ) 98 | → (x : Var) → Shift{ℓ}{V} k σ → σ x ≡ var→val (k + x) 99 | Shift-var σ zero x sft rewrite sft x = refl 100 | Shift-var {{S}} σ (suc k) x sft rewrite sft x 101 | | var→val-suc-shift {{S}} {x = k + x} 102 | | shifts-var→val {{S}} k x = refl 103 | 104 | up-seq : ∀{ℓ}{V₁ V₂ V₃ : Set ℓ}{{S : Shiftable V₁}} 105 | {{C : Composable V₁ V₂ V₃}} 106 | k (σ₂ : GSubst V₂) 107 | → ↑ k ⨟ σ₂ ≡ map-sub val₂₃ (drop k σ₂) 108 | up-seq {{S}}{{C}} k σ₂ = extensionality G 109 | where 110 | G : (x : Var) → (↑ k ⨟ σ₂) x ≡ map-sub (Composable.val₂₃ C) (drop k σ₂) x 111 | G x rewrite ⌈⌉-var→val σ₂ (k + x) = refl 112 | 113 | cons-seq : ∀{ℓ}{V₁ V₂ V₃ : Set ℓ} {{_ : Shiftable V₁}} {{_ : Shiftable V₃}} 114 | {{C : Composable V₁ V₂ V₃}} 115 | v₁ (σ₁ : GSubst V₁) (σ₂ : GSubst V₂) 116 | → (v₁ • σ₁) ⨟ σ₂ ≡ ⌈ σ₂ ⌉ v₁ • (σ₁ ⨟ σ₂) 117 | cons-seq {{C = C}} v₁ σ₁ σ₂ = extensionality G 118 | where 119 | G : (x : Var) → (v₁ • σ₁ ⨟ σ₂) x ≡ (Composable.⌈ C ⌉ σ₂ v₁ • (σ₁ ⨟ σ₂)) x 120 | G zero = refl 121 | G (suc x) = refl 122 | 123 | -------------------------------------------------------------------------------- /src/experimental/Fold.agda: -------------------------------------------------------------------------------- 1 | open import Data.Empty using (⊥) 2 | open import Data.Fin using (Fin; zero; suc; toℕ; inject≤; fromℕ<) 3 | open import Data.List using (List; []; _∷_; length; lookup; _++_) 4 | open import Data.List.Properties using (++-assoc; length-++) 5 | open import Data.Nat using (ℕ; zero; suc; _+_; _⊔_; _∸_; _<_; _≤_; z≤n; s≤s) 6 | open import Data.Nat.Properties 7 | using (+-suc; +-comm; ≤-step; ≤-refl; ≤-trans; m≤m+n; <⇒≤) 8 | open import Data.Product using (_×_) renaming (_,_ to ⟨_,_⟩ ) 9 | open import Data.Unit.Polymorphic using (⊤; tt) 10 | open import Environment 11 | open import Function using (_∘_) 12 | import Relation.Binary.PropositionalEquality as Eq 13 | open Eq using (_≡_; refl; sym; trans; cong; cong₂; cong-app) 14 | open Eq.≡-Reasoning 15 | open import Var 16 | open import ScopedTuple 17 | using (Tuple; map; _✖_; zip; zip-refl; map-pres-zip; map-compose-zip; 18 | map-compose; zip-map→rel; Lift-Eq-Tuple; Lift-Rel-Tuple; zip→rel) 19 | open import GenericSubstitution 20 | open import Agda.Primitive using (Level; lzero; lsuc) 21 | 22 | module experimental.Fold (Op : Set) (sig : Op → List ℕ) where 23 | 24 | open import AbstractBindingTree Op sig 25 | open import Map Op sig 26 | 27 | Bind : {ℓᶜ : Level} → Set → Set ℓᶜ → ℕ → Set ℓᶜ 28 | Bind V C zero = C 29 | Bind V C (suc b) = V → Bind V C b 30 | 31 | {------------------------------------------------------------------------------- 32 | Folding over an abstract binding tree 33 | ------------------------------------------------------------------------------} 34 | record Foldable {ℓᶜ : Level}(V : Set)(C : Set ℓᶜ) : Set (lsuc ℓᶜ) where 35 | field ret : V → C 36 | fold-op : (op : Op) → Tuple (sig op) (λ _ → C) → C 37 | binder : (op : Op) → (i : Fin (length (sig op))) 38 | → (j : Fin (lookup (sig op) i)) → Tuple (sig op) (Bind V C) → V 39 | 40 | open Env {{...}} 41 | open Foldable {{...}} 42 | 43 | arg-binder : ∀{ℓ}{V : Set}{C : Set ℓ} {{_ : Foldable V C}}{n : ℕ} 44 | → (op : Op) → (i : Fin (length (sig op))) → (b : Fin n) 45 | → (Bind V C (toℕ b)) → Tuple (sig op) (Bind V C) 46 | → { n≤ : n ≤ lookup (sig op) i} 47 | → C 48 | arg-binder op i zero r rs′ {n≤} = r 49 | arg-binder {n = suc n} op i (suc b) f rs′ {n≤} = 50 | let n≤′ = ≤-trans (≤-step ≤-refl) n≤ in 51 | let x = (f (binder op i (inject≤ b n≤′) rs′)) in 52 | arg-binder op i b x rs′ {n≤′} 53 | 54 | 55 | args-binder : ∀{bs}{ℓ}{V : Set}{C : Set ℓ} {{_ : Foldable V C}}{bs′} 56 | → (op : Op) → (i : Fin (length bs′)) 57 | → Tuple bs (Bind V C) 58 | → Tuple (sig op) (Bind V C) 59 | → { bs′++bs≡sig : bs′ ++ bs ≡ sig op } 60 | → Tuple bs (λ _ → C) 61 | args-binder {[]} op i tt rs′ {i+bs≡sig} = tt 62 | args-binder {b ∷ bs} {bs′ = bs′} op i ⟨ r , rs ⟩ rs′ {bs′++bs≡sig} = 63 | ⟨ arg-binder {n = lookup (sig op) (fromℕ< {length bs′}{length (sig op)} 64 | len[bs′] (var \ 'v) \ (var \ 'v)" (infixl "\" 75) where 18 | "(cons v \) 0 = v" | 19 | "(cons v \) (Suc x) = \ x" 20 | 21 | definition lift_ren :: "Renaming \ Renaming" ("\\<^sub>r") where 22 | "\\<^sub>r \ x \ Suc (\ x)" 23 | declare lift_ren_def[simp] 24 | 25 | locale subst1 = 26 | fixes Var :: "var \ 'a" 27 | and lift_sub :: "(var \ 'a) \ (var \ 'a)" ("\") 28 | and sub :: "(var \ 'a) \ 'a \ 'a" 29 | and ren :: "(var \ var) \ (var \ 'a)" 30 | and seq :: "(var \ 'a) \ (var \ 'a) \ (var \ 'a)" (infixr ";" 70) 31 | assumes ren_def: "ren \ x = Var (\ x)" 32 | and seq_def: "(\ ; \) x = sub \ (\ x)" 33 | and lift_sub_ren_var: "\ (ren \) x = Var (Suc (\ x))" 34 | begin 35 | 36 | lemma ren_cons[simp]: "ren (y \ \) = Var y \ ren \" 37 | apply (rule ext) 38 | apply (case_tac x) 39 | apply (simp add: ren_def) 40 | apply (simp add: ren_def) 41 | done 42 | 43 | lemma ren_shift[simp]: "ren (\\<^sub>r \) = \ (ren \)" 44 | apply (rule ext) 45 | apply (case_tac x) 46 | apply (simp add: lift_sub_ren_var ren_def lift_ren_def) 47 | apply (simp add: lift_ren_def ren_def lift_sub_ren_var) 48 | done 49 | 50 | theorem ConsSeq[simp]: "(M \ \) ; \ = sub \ M \ (\ ; \)" 51 | apply (rule ext) 52 | apply (case_tac x) 53 | apply (simp add: seq_def) 54 | apply (simp add: seq_def) 55 | done 56 | 57 | end 58 | 59 | locale subst2 = subst1 + 60 | assumes lift_def: "\ \ x = sub (ren Suc) (\ x)" 61 | and sub_var[simp]: "sub \ (Var x) = \ x" 62 | begin 63 | 64 | theorem shift_ren_seq[simp]: "\ (ren \ ; \) = \ (ren \) ; Var 0 \ \ \" 65 | apply (rule ext) 66 | apply (case_tac x) 67 | apply (simp add: seq_def ren_def lift_def sub_var) 68 | apply (simp add: seq_def ren_def lift_def sub_var) 69 | done 70 | 71 | lemma ren_suc: fixes \::Renaming shows "(ren \ ; ren Suc) = \ (ren \)" 72 | apply (rule ext) 73 | apply (simp add: seq_def ren_def lift_def) 74 | done 75 | 76 | lemma shift_seq_ren[simp]: "ren Suc ; Var 0 \ \ (ren \) = ren (\\<^sub>r \)" 77 | apply (rule ext) 78 | apply (case_tac x) 79 | apply (simp add: seq_def ren_def sub_var) 80 | apply (simp add: seq_def ren_def sub_var) 81 | done 82 | 83 | lemma shift_seq_sub[simp]: "ren Suc ; M \ \ = \" 84 | apply (rule ext) apply (case_tac x) 85 | apply (simp add: seq_def ren_def sub_var) 86 | apply (simp add: seq_def ren_def sub_var) 87 | done 88 | 89 | end 90 | 91 | locale subst3 = subst2 + 92 | assumes ren_sub: "sub \ (sub (ren \) M) = sub (ren \ ; \) M" 93 | begin 94 | 95 | lemma shift_sub_ren[simp]: "\ (\ ; ren \) = \ \ ; Var 0 \ ren (\\<^sub>r \)" 96 | apply (rule ext) 97 | apply (case_tac x) 98 | apply (simp add: seq_def ren_def ren_sub lift_def ren_suc) 99 | apply (simp add: seq_def ren_def ren_sub lift_def ren_suc) 100 | done 101 | 102 | end 103 | 104 | locale subst4 = subst3 + 105 | assumes sub_ren: "sub (ren \) (sub \ M) = sub (\ ; ren \) M" 106 | begin 107 | 108 | lemma sub_suc: fixes \::Renaming shows "(\ ; ren Suc) = \ \" 109 | unfolding seq_def apply (rule ext) 110 | apply (simp add: lift_def) 111 | done 112 | 113 | abbreviation exts :: "(var \ 'a) \ (var \ 'a)" where 114 | "exts \ \ Var 0 \ \ \" 115 | 116 | lemma exts_seq: "exts \ ; exts \ = exts (\ ; \)" 117 | apply (rule ext) apply (case_tac x) 118 | apply (simp add: seq_def lift_def sub_var) 119 | apply (simp add: seq_def lift_def ren_sub) 120 | using sub_ren sub_suc apply auto done 121 | 122 | end 123 | 124 | locale subst5 = subst4 + 125 | fixes ids :: "var \ 'a" 126 | assumes sub_sub[simp]: "sub \ (sub \ M) = sub (\ ; \) M" 127 | and ids_def: "ids x = Var x" 128 | begin 129 | 130 | lemma id_sub[simp]: "sub \ (ids x) = \ x" 131 | unfolding ids_def sub_var by simp 132 | 133 | lemma exts_id[simp]: "exts ids = ids" 134 | apply (rule ext) unfolding ids_def ren_def 135 | apply (case_tac x) apply force 136 | apply (simp add: ren_def lift_def sub_var) 137 | done 138 | 139 | theorem seq_ids_left[simp]: "ids ; \ = \" 140 | unfolding seq_def ids_def sub_var by simp 141 | 142 | theorem seq_assoc[simp]: "(\ ; \) ; \ = \ ; (\ ; \)" 143 | apply (rule ext) 144 | unfolding seq_def apply simp 145 | unfolding seq_def by simp 146 | 147 | lemma lift_sub_seq_suc[simp]: "\ \ = \ ; ren Suc" 148 | apply (rule ext) 149 | unfolding lift_def seq_def apply (simp add: ren_def) done 150 | 151 | lemma cons_shift_ids[simp]: "Var 0 \ ren Suc = ids" 152 | apply (rule ext) 153 | unfolding ids_def apply (case_tac x) apply simp apply (simp add: ren_def) done 154 | 155 | end 156 | 157 | locale subst6 = subst5 + 158 | fixes ext :: "(var \ 'a) \ (var \ 'a)" 159 | assumes sub_id[simp]: "sub ids M = M" 160 | and ext_def[simp]: "ext \ = Var 0 \ (\ ; ren Suc)" 161 | begin 162 | 163 | abbreviation subst_one :: "'a \ 'a \ 'a" ("_[_]" 70) where 164 | "subst_one N M \ sub (M \ ids) N" 165 | abbreviation subst_two :: "'a \ 'a \ 'a" ("_\_\" 60) where 166 | "N \ M \ \ sub (ext (M \ ids)) N" 167 | 168 | theorem seq_ids_right[simp]: "\ ; ids = \" 169 | unfolding seq_def by simp 170 | 171 | theorem subst_commute: "(sub (ext \) N) [ sub \ M ] = sub \ (N [ M ])" 172 | by simp 173 | 174 | theorem substitution: "M[N][L] = M\L\[N[L]]" 175 | by simp 176 | 177 | theorem ext_sub_cons: "(sub (ext \) N)[V] = sub (V \ \) N" 178 | by simp 179 | 180 | end 181 | 182 | 183 | end -------------------------------------------------------------------------------- /src/Map.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | open import Data.List using (List; []; _∷_) 3 | open import Data.Nat using (ℕ; zero; suc) 4 | open import Data.Product using (_×_) renaming (_,_ to ⟨_,_⟩ ) 5 | open import Data.Sum using (_⊎_; inj₁; inj₂) 6 | open import Structures 7 | open import Function using (_∘_) 8 | open import GSubst using (GSubst; ext) 9 | import Relation.Binary.PropositionalEquality as Eq 10 | open Eq using (_≡_; refl; sym; trans; cong; cong₂; cong-app; subst) 11 | open Eq.≡-Reasoning 12 | open import Sig 13 | open import Var 14 | 15 | module Map (Op : Set) (sig : Op → List Sig) where 16 | 17 | open import AbstractBindingTree Op sig 18 | 19 | map : ∀{ℓ}{V : Set ℓ} 20 | {{_ : Shiftable V}} {{_ : Quotable V}} 21 | → GSubst V → ABT → ABT 22 | 23 | map-arg : ∀{ℓ}{V : Set ℓ} 24 | {{_ : Shiftable V}} {{_ : Quotable V}} 25 | → GSubst V → {b : Sig} → Arg b → Arg b 26 | map-args : ∀{ℓ}{V : Set ℓ} 27 | {{_ : Shiftable V}} {{_ : Quotable V}} 28 | → GSubst V → {bs : List Sig} → Args bs → Args bs 29 | map σ (` x) = “ σ x ” 30 | map {V} σ (op ⦅ args ⦆) = op ⦅ map-args σ args ⦆ 31 | map-arg σ (ast M) = ast (map σ M) 32 | map-arg σ (bind M) = bind (map-arg (ext σ) M) 33 | map-arg σ (clear M) = clear M 34 | map-args σ {[]} nil = nil 35 | map-args σ {b ∷ bs} (cons x args) = cons (map-arg σ x) (map-args σ args) 36 | 37 | _○_≈_ : ∀{ℓ₁}{ℓ₂}{ℓ₃}{V₁ : Set ℓ₁}{V₂ : Set ℓ₂}{V₃ : Set ℓ₃} 38 | {{M₁ : Shiftable V₁}} {{M₂ : Shiftable V₂}} {{M₃ : Shiftable V₃}} 39 | {{Q₁ : Quotable V₁}}{{Q₂ : Quotable V₂}}{{Q₃ : Quotable V₃}} 40 | (σ₂ : GSubst V₂)(σ₁ : GSubst V₁)(σ₃ : GSubst V₃) → Set 41 | _○_≈_ {V₁}{V₂}{V₃}{{M₁}}{{M₂}}{{M₃}} σ₂ σ₁ σ₃ = 42 | ∀ x → map σ₂ “ σ₁ x ” ≡ “ σ₃ x ” 43 | 44 | map-map-fusion-ext : ∀{ℓ₁}{ℓ₂}{ℓ₃} {V₁ : Set ℓ₁} 45 | {V₂ : Set ℓ₂} {V₃ : Set ℓ₃} 46 | {{S₁ : Shiftable V₁}} {{S₂ : Shiftable V₂}} {{S₃ : Shiftable V₃}} 47 | {{Q₁ : Quotable V₁}} {{Q₂ : Quotable V₂}} {{Q₃ : Quotable V₃}} 48 | {σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{σ₃ : GSubst V₃} 49 | → (M : ABT) 50 | → σ₂ ○ σ₁ ≈ σ₃ 51 | → (∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{σ₃ : GSubst V₃} 52 | → σ₂ ○ σ₁ ≈ σ₃ → ext σ₂ ○ ext σ₁ ≈ ext σ₃) 53 | → map σ₂ (map σ₁ M) ≡ map σ₃ M 54 | map-map-fusion-ext (` x) σ₂○σ₁≈σ₃ mf-ext = σ₂○σ₁≈σ₃ x 55 | map-map-fusion-ext {V₁ = V₁}{V₂}{V₃} (op ⦅ args ⦆) σ₂○σ₁≈σ₃ mf-ext = 56 | cong (_⦅_⦆ op) (mmf-args args σ₂○σ₁≈σ₃) 57 | where 58 | mmf-arg : ∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{σ₃ : GSubst V₃}{b} (arg : Arg b) 59 | → σ₂ ○ σ₁ ≈ σ₃ 60 | → map-arg σ₂ (map-arg σ₁ arg) ≡ map-arg σ₃ arg 61 | mmf-args : ∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{σ₃ : GSubst V₃}{bs} 62 | (args : Args bs) 63 | → σ₂ ○ σ₁ ≈ σ₃ 64 | → map-args σ₂ (map-args σ₁ args) ≡ map-args σ₃ args 65 | mmf-arg (ast M) σ₂○σ₁≈σ₃ = 66 | cong ast (map-map-fusion-ext M σ₂○σ₁≈σ₃ mf-ext) 67 | mmf-arg (bind arg) σ₂○σ₁≈σ₃ = 68 | cong bind (mmf-arg arg (mf-ext σ₂○σ₁≈σ₃)) 69 | mmf-arg (clear arg) σ₂○σ₁≈σ₃ = refl 70 | mmf-args {bs = []} nil σ₂○σ₁≈σ₃ = refl 71 | mmf-args {bs = b ∷ bs} (cons arg args) σ₂○σ₁≈σ₃ = 72 | cong₂ cons (mmf-arg arg σ₂○σ₁≈σ₃) (mmf-args args σ₂○σ₁≈σ₃) 73 | 74 | _≃_ : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 75 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 76 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 77 | (σ₁ : GSubst V₁)(σ₂ : GSubst V₂) → Set 78 | _≃_ σ₁ σ₂ = ∀ x → “ σ₁ x ” ≡ “ σ₂ x ” 79 | 80 | {- todo: generalize to map-cong to simulation -} 81 | map-cong : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 82 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 83 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 84 | {σ₁ : GSubst V₁}{σ₂ : GSubst V₂} 85 | → (M : ABT) 86 | → σ₁ ≃ σ₂ 87 | → (∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂} → σ₁ ≃ σ₂ → ext σ₁ ≃ ext σ₂) 88 | → map σ₁ M ≡ map σ₂ M 89 | map-cong (` x) σ₁≃σ₂ mc-ext = σ₁≃σ₂ x 90 | map-cong {ℓ}{V₁}{V₂} (op ⦅ args ⦆) σ₁≃σ₂ mc-ext = 91 | cong (_⦅_⦆ op) (mc-args args σ₁≃σ₂) 92 | where 93 | mc-arg : ∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{b} (arg : Arg b) → σ₁ ≃ σ₂ 94 | → map-arg σ₁ arg ≡ map-arg σ₂ arg 95 | mc-args : ∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{bs} (args : Args bs) → σ₁ ≃ σ₂ 96 | → map-args σ₁ args ≡ map-args σ₂ args 97 | mc-arg (ast M) σ₁≃σ₂ = 98 | cong ast (map-cong M σ₁≃σ₂ mc-ext) 99 | mc-arg (bind arg) σ₁≃σ₂ = 100 | cong bind (mc-arg arg (mc-ext σ₁≃σ₂)) 101 | mc-arg (clear arg) σ₁≃σ₂ = refl 102 | mc-args {bs = []} nil σ₁≃σ₂ = refl 103 | mc-args {σ₁}{σ₂} {b ∷ bs} (cons arg args) σ₁≃σ₂ = 104 | cong₂ cons (mc-arg arg σ₁≃σ₂) (mc-args args σ₁≃σ₂) 105 | 106 | _⊢_≃_ : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 107 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 108 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 109 | (M : ABT)(σ₂ : GSubst V₂)(σ₁ : GSubst V₁) → Set 110 | _⊢_≃_ M σ₁ σ₂ = ∀ x → FV M x → “ σ₁ x ” ≡ “ σ₂ x ” 111 | 112 | _⊢ₐ_≃_ : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 113 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 114 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 115 | {b : Sig}(arg : Arg b)(σ₂ : GSubst V₂)(σ₁ : GSubst V₁) → Set 116 | _⊢ₐ_≃_ {b} arg σ₁ σ₂ = ∀ x → FV-arg arg x → “ σ₁ x ” ≡ “ σ₂ x ” 117 | 118 | _⊢₊_≃_ : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 119 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 120 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 121 | {bs : List Sig}(args : Args bs)(σ₂ : GSubst V₂)(σ₁ : GSubst V₁) → Set 122 | _⊢₊_≃_ {bs} args σ₁ σ₂ = ∀ x → FV-args args x → “ σ₁ x ” ≡ “ σ₂ x ” 123 | 124 | 125 | map-cong-FV : ∀{ℓ}{V₁ : Set ℓ}{V₂ : Set ℓ} 126 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 127 | {{_ : Quotable V₁}} {{_ : Quotable V₂}} 128 | {σ₁ : GSubst V₁}{σ₂ : GSubst V₂} 129 | → (M : ABT) 130 | → M ⊢ σ₁ ≃ σ₂ 131 | → (∀{b}{arg : Arg b}{σ₁ : GSubst V₁}{σ₂ : GSubst V₂} 132 | → bind arg ⊢ₐ σ₁ ≃ σ₂ → arg ⊢ₐ ext σ₁ ≃ ext σ₂) 133 | → map σ₁ M ≡ map σ₂ M 134 | map-cong-FV (` x) σ₁≃σ₂ mc-ext = σ₁≃σ₂ x refl 135 | map-cong-FV {V₁ = V₁}{V₂}(op ⦅ args ⦆) σ₁≃σ₂ mc-ext = 136 | cong (_⦅_⦆ op) (mc-args args σ₁≃σ₂) 137 | where 138 | mc-arg : ∀{σ₁ : GSubst V₁}{ σ₂ : GSubst V₂}{b} (arg : Arg b) 139 | → arg ⊢ₐ σ₁ ≃ σ₂ 140 | → map-arg σ₁ arg ≡ map-arg σ₂ arg 141 | mc-args : ∀{σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{bs} (args : Args bs) 142 | → args ⊢₊ σ₁ ≃ σ₂ 143 | → map-args σ₁ args ≡ map-args σ₂ args 144 | mc-arg (ast M) σ₁≃σ₂ = 145 | cong ast (map-cong-FV M σ₁≃σ₂ (λ{b}{arg} → mc-ext {b}{arg})) 146 | mc-arg {σ₁}{σ₂}{b = ν b} (bind arg) σ₁≃σ₂ = 147 | cong bind (mc-arg arg (mc-ext {b}{arg} σ₁≃σ₂)) 148 | mc-arg {σ₁}{σ₂} (clear arg) σ₁≃σ₂ = refl 149 | mc-args {bs = []} nil σ₁≃σ₂ = refl 150 | mc-args {σ₁}{σ₂}{bs = b ∷ bs} (cons arg args) σ₁≃σ₂ = 151 | cong₂ cons (mc-arg arg G) (mc-args args H) 152 | where 153 | G : arg ⊢ₐ σ₁ ≃ σ₂ 154 | G x x∈arg = σ₁≃σ₂ x (inj₁ x∈arg) 155 | H : args ⊢₊ σ₁ ≃ σ₂ 156 | H x x∈args = σ₁≃σ₂ x (inj₂ x∈args) 157 | 158 | -------------------------------------------------------------------------------- /src/ScopedTuple.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Data.List using (List; []; _∷_) 3 | open import Data.Nat using (ℕ; zero; suc; _+_; _∸_) 4 | open import Data.Product using (_×_; proj₁; proj₂) renaming (_,_ to ⟨_,_⟩ ) 5 | open import Data.Unit.Polymorphic using (⊤; tt) 6 | open import Data.Unit renaming (⊤ to Unit; tt to unit) 7 | open import Function using (_∘_) 8 | open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; cong₂) 9 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 10 | open import Sig 11 | 12 | module ScopedTuple where 13 | 14 | {- Scet: A scoped Set -} 15 | Scet : {ℓ : Level} → Set (lsuc ℓ) 16 | Scet {ℓ} = Sig → Set ℓ 17 | 18 | _⇨_ : {ℓ : Level} → Scet {ℓ} → Scet {ℓ} → Set ℓ 19 | A ⇨ B = (∀ {b : Sig} → A b → B b) 20 | 21 | 𝒫 : {ℓ : Level} → Scet {ℓ} → Set (lsuc ℓ) 22 | 𝒫 {ℓ} A = (∀ {b : Sig} → A b → Set ℓ) 23 | 24 | _✖_ : {ℓ : Level} → Scet {ℓ} → Scet {ℓ} → Set (lsuc ℓ) 25 | _✖_ {ℓ} A B = (∀ {b : Sig} → A b → B b → Set ℓ) 26 | 27 | Sigs : Set 28 | Sigs = List Sig 29 | 30 | Tuple : {ℓ : Level} → Sigs → Scet {ℓ} → Set ℓ 31 | Tuple [] A = ⊤ 32 | Tuple (b ∷ bs) A = A b × Tuple bs A 33 | 34 | map : ∀{ℓ : Level}{A : Scet {ℓ}}{B : Scet {ℓ}} → (A ⇨ B) → {bs : Sigs} 35 | → Tuple {ℓ} bs A → Tuple {ℓ} bs B 36 | map f {[]} ⊤ = tt 37 | map f {b ∷ bs} ⟨ x , xs ⟩ = ⟨ f x , map f xs ⟩ 38 | 39 | foldr : ∀{ℓ : Level}{A}{B : Set} → (∀ {b} → A b → B → B) 40 | → B → {bs : Sigs} → Tuple {ℓ} bs A → B 41 | foldr c n {[]} tt = n 42 | foldr c n {b ∷ bs} ⟨ x , xs ⟩ = c x (foldr c n xs) 43 | 44 | all : ∀{A} → 𝒫 A → {bs : Sigs} → Tuple bs A → Set 45 | all {A} P {[]} tt = ⊤ 46 | all {A} P {b ∷ bs} ⟨ x , xs ⟩ = P x × (all P xs) 47 | 48 | zip : ∀{ℓ}{A B : Scet {ℓ}} → A ✖ B → {bs : Sigs} 49 | → Tuple bs A → Tuple bs B → Set ℓ 50 | zip R {[]} tt tt = ⊤ 51 | zip R {b ∷ bs} ⟨ a₁ , as₁ ⟩ ⟨ a₂ , as₂ ⟩ = R a₁ a₂ × zip R as₁ as₂ 52 | 53 | map-cong : ∀{ℓ}{A B : Scet {ℓ}}{f g : A ⇨ B} {bs} {xs : Tuple bs A} 54 | → (∀{b} (x : A b) → f x ≡ g x) 55 | → map f xs ≡ map g xs 56 | map-cong {bs = []} {tt} eq = refl 57 | map-cong {bs = b ∷ bs} {⟨ x , xs ⟩} eq = cong₂ ⟨_,_⟩ (eq x) (map-cong eq) 58 | 59 | map-compose : ∀{ℓ}{A B C : Scet {ℓ}} {g : B ⇨ C} {f : A ⇨ B} {bs : Sigs} 60 | {xs : Tuple bs A} 61 | → (map g (map f xs)) ≡ (map (g ∘ f) xs) 62 | map-compose {A = A}{B}{C} {g} {f} {[]} {tt} = refl 63 | map-compose {A = A}{B}{C} {g} {f} {b ∷ bs} {⟨ x , xs ⟩} = 64 | cong₂ ⟨_,_⟩ refl map-compose 65 | 66 | tuple-pred : ∀{ℓ}{A : Scet {ℓ}}{P : 𝒫 A} 67 | → (P× : ∀ bs → Tuple bs A → Set) 68 | → (∀ (b : Sig) → (a : A b) → P {b} a) 69 | → {bs : Sigs} → (xs : Tuple bs A) 70 | → (P× [] tt) 71 | → (∀{b : Sig}{bs : Sigs}{x xs} 72 | → P {b} x → P× bs xs → P× (b ∷ bs) ⟨ x , xs ⟩) 73 | → P× bs xs 74 | tuple-pred {A} {P} P× f {[]} tt base step = base 75 | tuple-pred {A} {P} P× f {x ∷ bs} ⟨ fst , snd ⟩ base step = 76 | step (f x fst) (tuple-pred P× f snd base step) 77 | 78 | 79 | all-intro : ∀{A : Scet} → (P : 𝒫 A) 80 | → (∀ {b} (a : A b) → P {b} a) 81 | → {bs : Sigs} → (xs : Tuple bs A) 82 | → all P xs 83 | all-intro {A} P f {[]} tt = tt 84 | all-intro {A} P f {b ∷ bs} ⟨ x , xs ⟩ = ⟨ (f x) , (all-intro P f xs) ⟩ 85 | 86 | zip-refl : ∀{ℓ}{bs A} (xs : Tuple {ℓ} bs A) → zip {ℓ} _≡_ xs xs 87 | zip-refl {ℓ}{[]} tt = tt 88 | zip-refl {ℓ}{b ∷ bs} {A} ⟨ x , xs ⟩ = ⟨ refl , zip-refl xs ⟩ 89 | 90 | zip-intro : ∀{ℓ}{A B : Scet {ℓ}} → (R : A ✖ B) 91 | → (∀ {c} (a : A c) (b : B c) → R {c} a b) 92 | → {bs : Sigs} → (xs : Tuple bs A) → (ys : Tuple bs B) 93 | → zip R xs ys 94 | zip-intro {A} {B} R f {[]} tt tt = tt 95 | zip-intro {A} {B} R f {b ∷ bs} ⟨ x , xs ⟩ ⟨ y , ys ⟩ = 96 | ⟨ (f x y) , (zip-intro R f xs ys) ⟩ 97 | 98 | map-pres-zip : ∀{ℓ}{bs}{A1 B1 : Scet {ℓ}}{A2 B2 : Scet {ℓ}} {xs ys} 99 | → (P : A1 ✖ B1) → (Q : A2 ✖ B2) → (f : A1 ⇨ A2) → (g : B1 ⇨ B2) 100 | → zip (λ{b} → P {b}) {bs} xs ys 101 | → (∀{b}{x}{y} → P {b} x y → Q (f x) (g y)) 102 | → zip Q (map f xs) (map g ys) 103 | map-pres-zip {ℓ}{bs = []} P Q f g tt pres = tt 104 | map-pres-zip {bs = b ∷ bs}{xs = ⟨ x , xs ⟩} {⟨ y , ys ⟩} P Q f g ⟨ z , zs ⟩ 105 | pres = 106 | ⟨ pres z , map-pres-zip P Q f g zs pres ⟩ 107 | 108 | record Lift-Pred-Tuple {ℓ}{A : Scet{ℓ}} (P : 𝒫 A) 109 | (P× : ∀{bs} → Tuple bs A → Set) : Set ℓ where 110 | field base : (P× {bs = []} tt) 111 | step : (∀{b : Sig}{bs : Sigs}{x xs} 112 | → P {b} x → P× {bs} xs → P× ⟨ x , xs ⟩) 113 | 114 | record Lift-Rel-Tuple {ℓ}{A B : Scet{ℓ}} (R : A ✖ B) 115 | (R× : ∀{bs} → Tuple bs A → Tuple bs B → Set) : Set ℓ where 116 | field base : (R× {bs = []} tt tt) 117 | step : (∀{b : Sig}{bs : Sigs}{x xs}{y ys} 118 | → R {b} x y → R× {bs} xs ys → R× ⟨ x , xs ⟩ ⟨ y , ys ⟩) 119 | 120 | Lift-Eq-Tuple : ∀{A : Set} → Lift-Rel-Tuple {A = λ _ → A}{λ _ → A} _≡_ _≡_ 121 | Lift-Eq-Tuple = record { base = refl ; step = λ { refl refl → refl } } 122 | 123 | all→pred : ∀{bs A xs} 124 | → (P : 𝒫 A) → (P× : ∀ {bs} → Tuple bs A → Set) 125 | → (L : Lift-Pred-Tuple P P×) 126 | → all P {bs} xs → P× xs 127 | all→pred {[]} {xs = tt} P P× L tt = Lift-Pred-Tuple.base L 128 | all→pred {b ∷ bs} {xs = ⟨ x , xs ⟩} P P× L ⟨ z , zs ⟩ = 129 | let IH = all→pred {bs} {xs = xs} P P× L zs in 130 | Lift-Pred-Tuple.step L z IH 131 | 132 | lift-pred : ∀{A : Scet} → (P : 𝒫 A) → (P× : ∀ {bs} → Tuple bs A → Set) 133 | → (L : Lift-Pred-Tuple P P×) 134 | → (∀ {b} (a : A b) → P {b} a) 135 | → {bs : Sigs} → (xs : Tuple bs A) 136 | → P× xs 137 | lift-pred {A} P P× L f {bs} xs = 138 | all→pred {bs}{A}{xs} P P× L (all-intro {A} P f {bs} xs) 139 | 140 | zip→rel : ∀{ℓ}{bs}{A B : Scet{ℓ}}{xs ys} 141 | → (R : A ✖ B) → (R× : ∀ {bs} → Tuple bs A → Tuple bs B → Set) 142 | → (L : Lift-Rel-Tuple R R×) 143 | → zip R {bs} xs ys → R× xs ys 144 | zip→rel {bs = []} {xs = tt} {tt} R R× L tt = Lift-Rel-Tuple.base L 145 | zip→rel {bs = b ∷ bs} {xs = ⟨ x , xs ⟩} {⟨ y , ys ⟩} R R× L ⟨ z , zs ⟩ = 146 | let IH = zip→rel {bs = bs} {xs = xs} {ys} R R× L zs in 147 | Lift-Rel-Tuple.step L z IH 148 | 149 | zip-map→rel : ∀{ℓ}{bs}{A1 B1 : Scet {ℓ}}{A2 B2 : Scet {ℓ}}{xs ys} 150 | → (P : A1 ✖ B1) → (Q : A2 ✖ B2) 151 | → (R : ∀ {bs} → Tuple bs A2 → Tuple bs B2 → Set) 152 | → (f : A1 ⇨ A2) → (g : B1 ⇨ B2) 153 | → (∀{b}{x}{y} → P{b} x y → Q{b} (f x) (g y)) 154 | → (L : Lift-Rel-Tuple Q R) 155 | → zip P {bs} xs ys → R {bs} (map f xs) (map g ys) 156 | zip-map→rel P Q R f g P→Q L zs = zip→rel Q R L (map-pres-zip P Q f g zs P→Q) 157 | 158 | map-compose-zip : ∀{ℓ}{A B C C′ : Scet{ℓ}} 159 | {g : B ⇨ C} {f : A ⇨ B}{h : A ⇨ C′} 160 | {bs : Sigs}{R : C ✖ C′} 161 | {xs : Tuple bs A} 162 | → (∀ {b : Sig} x → R {b} (g (f x)) (h x)) 163 | → zip R (map g (map f xs)) (map h xs) 164 | map-compose-zip {bs = []} gf=h = tt 165 | map-compose-zip {bs = b ∷ bs} {R} {⟨ x , xs ⟩} gf=h = 166 | ⟨ (gf=h x) , (map-compose-zip gf=h) ⟩ 167 | 168 | 169 | -------------------------------------------------------------------------------- /src/Fold.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K #-} 2 | 3 | {- 4 | 5 | This fold is going to be replaced by Fold2 once that version catches 6 | up to this one. 7 | 8 | -} 9 | 10 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) 11 | open import Data.Empty using (⊥) 12 | open import Data.List using (List; []; _∷_) renaming (map to lmap) 13 | open import Data.Nat using (ℕ; zero; suc; _+_; _∸_) 14 | open import Data.Product 15 | using (_×_; proj₁; proj₂; Σ-syntax) renaming (_,_ to ⟨_,_⟩ ) 16 | open import Data.Sum using (_⊎_; inj₁; inj₂) 17 | open import Data.Unit.Polymorphic using (⊤; tt) 18 | open import Level using (levelOfType) 19 | open import Structures 20 | open import Function using (_∘_) 21 | import Relation.Binary.PropositionalEquality as Eq 22 | open Eq using (_≡_; refl; sym; cong; cong₂; cong-app) 23 | open Eq.≡-Reasoning 24 | open import Var 25 | open import ScopedTuple 26 | using (Tuple; map; _✖_; zip; zip-refl; map-pres-zip; map-compose-zip; 27 | map-compose; zip-map→rel; Lift-Eq-Tuple; Lift-Rel-Tuple; zip→rel) 28 | open import GSubst 29 | open import GenericSubstitution 30 | open import Sig 31 | 32 | module Fold (Op : Set) (sig : Op → List Sig) where 33 | 34 | open import AbstractBindingTree Op sig 35 | open Structures.WithOpSig Op sig 36 | 37 | private 38 | variable 39 | ℓ : Level 40 | V V₁ V₂ C C₁ C₂ : Set ℓ 41 | 42 | {------------------------------------------------------------------------------- 43 | Folding over an abstract binding tree 44 | ------------------------------------------------------------------------------} 45 | 46 | fold : {{_ : Shiftable V}} {{_ : Foldable V C}} 47 | → GSubst V → ABT → C 48 | fold-arg : {{_ : Shiftable V}} {{_ : Foldable V C}} 49 | → GSubst V → {b : Sig} → Arg b → Bind V C b 50 | fold-args : {{_ : Shiftable V}} {{_ : Foldable V C}} 51 | → GSubst V → {bs : List Sig} → Args bs → Tuple bs (Bind V C) 52 | 53 | fold σ (` x) = ret (σ x) 54 | fold σ (op ⦅ args ⦆) = fold-op op (fold-args σ {sig op} args) 55 | fold-arg σ (ast M) = (fold σ M) 56 | fold-arg σ (bind arg) v = fold-arg (v • ⟰ σ) arg 57 | fold-arg σ (clear arg) = fold-arg id arg 58 | fold-args σ {[]} nil = tt 59 | fold-args σ {b ∷ bs} (cons arg args) = ⟨ fold-arg σ arg , fold-args σ args ⟩ 60 | 61 | {------------------------------------------------------------------------------- 62 | Simulation between two folds 63 | ------------------------------------------------------------------------------} 64 | 65 | _≅_ : {ℓ : Level}{V₁ V₂ : Set ℓ} {{_ : Equiv V₁ V₂}} 66 | (σ₁ : GSubst V₁) (σ₂ : GSubst V₂) → Set ℓ 67 | _≅_ σ₁ σ₂ = ∀ x → σ₁ x ≈ σ₂ x 68 | 69 | sim-ext : {σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{v₁ : V₁}{v₂ : V₂} 70 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 71 | {{_ : Relatable V₁ V₂}} 72 | → σ₁ ≅ σ₂ → v₁ ≈ v₂ → (v₁ • ⟰ σ₁) ≅ (v₂ • ⟰ σ₂) 73 | sim-ext {σ₁} {σ₂} {v₁} {v₂} σ₁≅σ₂ v₁≈v₂ zero = v₁≈v₂ 74 | sim-ext {σ₁} {σ₂} {v₁} {v₂} σ₁≅σ₂ v₁≈v₂ (suc x) = shift≈ (σ₁≅σ₂ x) 75 | 76 | sim : ∀ {σ₁ : GSubst V₁}{σ₂ : GSubst V₂} 77 | {{S1 : Shiftable V₁}} {{S2 : Shiftable V₂}} 78 | {{F1 : Foldable V₁ C₁}} {{F2 : Foldable V₂ C₂}} 79 | {{EqV : Equiv V₁ V₂}} {{EqC : Equiv C₁ C₂}} {{Sim : Similar V₁ V₂ C₁ C₂}} 80 | → (M : ABT) 81 | → σ₁ ≅ σ₂ 82 | → (fold σ₁ M) ≈ (fold σ₂ M) 83 | 84 | sim-arg : ∀ {σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{b} (arg : Arg b) 85 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 86 | {{_ : Foldable V₁ C₁}} {{_ : Foldable V₂ C₂}} 87 | {{_ : Equiv C₁ C₂}} {{_ : Similar V₁ V₂ C₁ C₂}} 88 | → σ₁ ≅ σ₂ → (_⩳_ {b = b}) (fold-arg σ₁ {b} arg) (fold-arg σ₂ {b} arg) 89 | 90 | sim-args : ∀ {σ₁ : GSubst V₁}{σ₂ : GSubst V₂}{bs} (args : Args bs) 91 | {{_ : Shiftable V₁}} {{_ : Shiftable V₂}} 92 | {{_ : Foldable V₁ C₁}} {{_ : Foldable V₂ C₂}} 93 | {{_ : Equiv C₁ C₂}} {{_ : Similar V₁ V₂ C₁ C₂}} 94 | → σ₁ ≅ σ₂ 95 | → zip (λ {b} → _⩳_{V₁ = V₁}{V₂}{C₁}{C₂}{b = b}) (fold-args σ₁ {bs} args) 96 | (fold-args σ₂ {bs} args) 97 | 98 | sim (` x) σ₁≅σ₂ = ret≈ (σ₁≅σ₂ x) 99 | sim {V₁ = V₁}{V₂}{C₁}{C₂}{σ₁}{σ₂} (op ⦅ args ⦆) σ₁≅σ₂ = 100 | op⩳ (sim-args args σ₁≅σ₂) 101 | 102 | sim-arg (ast M) σ₁≊σ₂ = sim M σ₁≊σ₂ 103 | sim-arg {b = ν b} (bind arg) σ₁≊σ₂ v₁≈v₂ = 104 | sim-arg {b = b} arg (sim-ext σ₁≊σ₂ v₁≈v₂) 105 | sim-arg (clear arg) σ₁≊σ₂ = sim-arg arg λ x → var→val≈ x 106 | 107 | sim-args {bs = []} args σ₁≊σ₂ = tt 108 | sim-args {bs = b ∷ bs} (cons arg args) σ₁≊σ₂ = 109 | ⟨ sim-arg arg σ₁≊σ₂ , sim-args args σ₁≊σ₂ ⟩ 110 | 111 | 112 | fold-refl : ∀ {σ : GSubst V} 113 | {{_ : Shiftable V}} {{_ : Foldable V C}} 114 | {{_ : Equiv C C}} {{_ : Similar V V C C}} 115 | → (M : ABT) 116 | → σ ≅ σ 117 | → fold σ M ≈ fold σ M 118 | fold-refl M σ≅σ = sim M σ≅σ 119 | 120 | fold-arg-refl : ∀ {σ : GSubst V}{b} (arg : Arg b) 121 | {{_ : Shiftable V}} {{_ : Foldable V C}} 122 | {{_ : Equiv C C}} {{_ : Similar V V C C}} 123 | → σ ≅ σ → (_⩳_ {b = b}) (fold-arg σ {b} arg) (fold-arg σ {b} arg) 124 | fold-arg-refl arg σ≅σ = sim-arg arg σ≅σ 125 | 126 | 127 | {------------------------------------------------------------------------------- 128 | FV of fold 129 | ------------------------------------------------------------------------------} 130 | 131 | fv-env : {{_ : SyntacticFold V C}} → GSubst V → Var → Set 132 | fv-env γ x = Σ[ y ∈ Var ] fvᵛ (γ y) x 133 | 134 | 135 | fv-bind : {{_ : SyntacticFold V C}} {b : Sig} → Bind V C b → Var → Set 136 | fv-bind {b = ■} r x = fvᶜ r x 137 | fv-bind {b = ν b} r x = fv-bind {b = b} (r (var→val 0)) (suc x) 138 | fv-bind {b = ∁ b} r x = ⊥ 139 | 140 | fv-binds : {{_ : SyntacticFold V C}} {bs : List Sig} 141 | → Tuple bs (Bind V C) → Var → Set 142 | fv-binds {bs = []} tt x = ⊥ 143 | fv-binds {bs = b ∷ bs} ⟨ r , rs ⟩ x = fv-bind {b = b} r x ⊎ fv-binds rs x 144 | 145 | FV-fold : {{_ : SyntacticFold V C}} 146 | (γ : GSubst V) (M : ABT) (x : Var) 147 | → ((γ : GSubst V) (op : Op) (args : Args (sig op)) (x : Var) 148 | → fvᶜ (fold-op op (fold-args γ args)) x 149 | → fv-binds (fold-args γ args) x) 150 | → fvᶜ (fold γ M) x 151 | → fv-env γ x 152 | 153 | FV-fold γ (` y) x fv-op fv-fold rewrite fv-ret (γ y) = ⟨ y , fv-fold ⟩ 154 | FV-fold {V = V}{C} γ (op ⦅ args ⦆) x fv-op fv-fold = 155 | FV-fold-args γ args x (fv-op γ op args x fv-fold) 156 | where 157 | FV-fold-arg : ∀ (γ : GSubst V) {b : Sig} (arg : Arg b) (x : Var) 158 | → fv-bind {b = b} (fold-arg γ arg) x → fv-env γ x 159 | FV-fold-arg γ (ast M) x fv-fold = FV-fold γ M x fv-op fv-fold 160 | FV-fold-arg γ (bind arg) x fv-fold 161 | with FV-fold-arg ((var→val 0) • ⟰ γ) arg (suc x) fv-fold 162 | ... | ⟨ suc y , fvγ'y ⟩ = ⟨ y , fv-shift (γ y) x fvγ'y ⟩ 163 | ... | ⟨ 0 , fvγ'y ⟩ rewrite fv-var→val {V = V} 0 (suc x) 164 | with fvγ'y 165 | ... | () 166 | FV-fold-arg γ (clear arg) x () 167 | 168 | FV-fold-args : ∀ (γ : GSubst V) {bs : List Sig} (args : Args bs) (x : Var) 169 | → fv-binds (fold-args γ args) x → fv-env γ x 170 | FV-fold-args γ nil x () 171 | FV-fold-args γ (cons arg args) x (inj₁ fv-fld) = FV-fold-arg γ arg x fv-fld 172 | FV-fold-args γ (cons arg args) x (inj₂ fv-fld) = FV-fold-args γ args x fv-fld 173 | -------------------------------------------------------------------------------- /src/rewriting/examples/Experiments.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --rewriting #-} 2 | module rewriting.examples.Experiments where 3 | 4 | open import Data.List using (List; []; _∷_; length; map) 5 | open import Data.Nat 6 | open import Data.Bool using (true; false) renaming (Bool to 𝔹) 7 | open import Data.Nat.Properties 8 | open import Data.Product using (_,_;_×_; proj₁; proj₂; Σ-syntax; ∃-syntax) 9 | open import Data.Unit using (⊤; tt) 10 | open import Data.Unit.Polymorphic renaming (⊤ to topᵖ; tt to ttᵖ) 11 | open import Data.Empty using (⊥; ⊥-elim) 12 | open import Data.Sum using (_⊎_; inj₁; inj₂) 13 | open import Relation.Binary.PropositionalEquality as Eq 14 | using (_≡_; _≢_; refl; sym; cong; subst; trans) 15 | open import Relation.Nullary using (¬_; Dec; yes; no) 16 | open import Var 17 | open import rewriting.examples.Cast 18 | open import rewriting.examples.CastDeterministic 19 | open import rewriting.examples.StepIndexedLogic2 20 | 21 | ℰ⊎𝒱-type : Set 22 | ℰ⊎𝒱-type = (Prec × Term × Term) ⊎ (Prec × Term × Term) 23 | 24 | ℰ⊎𝒱-ctx : Context 25 | ℰ⊎𝒱-ctx = ℰ⊎𝒱-type ∷ [] 26 | 27 | ℰˢ⟦_⟧ : Prec → Term → Term → Setˢ ℰ⊎𝒱-ctx (cons Now ∅) 28 | ℰˢ⟦ A⊑B ⟧ M M′ = (inj₂ (A⊑B , M , M′)) ∈ zeroˢ 29 | 30 | 𝒱ˢ⟦_⟧ : Prec → Term → Term → Setˢ ℰ⊎𝒱-ctx (cons Now ∅) 31 | 𝒱ˢ⟦ A⊑B ⟧ V V′ = (inj₁ (A⊑B , V , V′)) ∈ zeroˢ 32 | 33 | pre-𝒱 : Prec → Term → Term → Setˢ ℰ⊎𝒱-ctx (cons Later ∅) 34 | pre-𝒱 (.★ , .★ , unk⊑unk) (V ⟨ G !⟩) (V′ ⟨ H !⟩) 35 | with G ≡ᵍ H 36 | ... | yes refl = let g = gnd⇒ty G in 37 | (Value V)ˢ ×ˢ (Value V′)ˢ 38 | ×ˢ (▷ˢ (𝒱ˢ⟦ (g , g , Refl⊑) ⟧ V V′)) 39 | ... | no neq = ⊥ ˢ 40 | pre-𝒱 (.★ , .★ , unk⊑unk) V V′ = ⊥ ˢ 41 | pre-𝒱 (.★ , .A′ , unk⊑any{G}{A′} G⊑A′) (V ⟨ H !⟩) V′ 42 | with G ≡ᵍ H 43 | ... | yes refl = (Value V)ˢ ×ˢ (Value V′)ˢ 44 | ×ˢ ▷ˢ (𝒱ˢ⟦ (gnd⇒ty G , A′ , G⊑A′) ⟧ V V′) 45 | ... | no new = ⊥ ˢ 46 | pre-𝒱 (.★ , .A′ , unk⊑any{G}{A′} G⊑A′) V V′ = ⊥ ˢ 47 | pre-𝒱 ($ₜ ι , $ₜ ι , base⊑) ($ c) ($ c′) = (c ≡ c′) ˢ 48 | pre-𝒱 ($ₜ ι , $ₜ ι , base⊑) V V′ = ⊥ ˢ 49 | pre-𝒱 ((A ⇒ B) , (A′ ⇒ B′) , fun⊑ A⊑A′ B⊑B′) (ƛ N) (ƛ N′) = 50 | ∀ˢ[ W ] ∀ˢ[ W′ ] ▷ˢ (𝒱ˢ⟦ (A , A′ , A⊑A′) ⟧ W W′) 51 | →ˢ ▷ˢ (ℰˢ⟦ (B , B′ , B⊑B′) ⟧ (N [ W ]) (N′ [ W′ ])) 52 | pre-𝒱 ((A ⇒ B) , (A′ ⇒ B′) , fun⊑ A⊑A′ B⊑B′) V V′ = ⊥ ˢ 53 | 54 | instance 55 | TermInhabited : Inhabited Term 56 | TermInhabited = record { elt = ` 0 } 57 | 58 | {- Right-to-left version -} 59 | pre-ℰ : Prec → Term → Term → Setˢ ℰ⊎𝒱-ctx (cons Later ∅) 60 | pre-ℰ c M M′ = 61 | ((Value M)ˢ ×ˢ (Value M′)ˢ × (pre-𝒱 c M M′)) 62 | ⊎ˢ ((Value M′)ˢ ×ˢ (∃ˢ[ N ] (M —→ N)ˢ ×ˢ ▷ˢ ℰˢ⟦ c ⟧ V M′)) 63 | ⊎ˢ (∃ˢ[ N′ ] (M′ —→ N′)ˢ ×ˢ ▷ˢ (ℰˢ⟦ c ⟧ M N′)) 64 | ⊎ˢ (Blame M′)ˢ 65 | 66 | pre-ℰ⊎𝒱 : ℰ⊎𝒱-type → Setˢ ℰ⊎𝒱-ctx (cons Later ∅) 67 | pre-ℰ⊎𝒱 (inj₁ (c , V , V′)) = pre-𝒱 c V V′ 68 | pre-ℰ⊎𝒱 (inj₂ (c , M , M′)) = pre-ℰ c M M′ 69 | 70 | ℰ⊎𝒱 : ℰ⊎𝒱-type → Setᵒ 71 | ℰ⊎𝒱 X = μᵒ pre-ℰ⊎𝒱 X 72 | 73 | abstract 74 | 𝒱⟦_⟧ : (c : Prec) → Term → Term → Setᵒ 75 | 𝒱⟦ c ⟧ V V′ = ℰ⊎𝒱 (inj₁ (c , V , V′)) 76 | 77 | ℰ⟦_⟧ : (c : Prec) → Term → Term → Setᵒ 78 | ℰ⟦ c ⟧ M M′ = ℰ⊎𝒱 (inj₂ (c , M , M′)) 79 | 80 | ℰ-stmt : ∀{c}{M M′} 81 | → ℰ⟦ c ⟧ M M′ ≡ᵒ 82 | (((Value M′)ᵒ ×ᵒ (∃ᵒ[ V ] (M —↠ V)ᵒ ×ᵒ (Value V)ᵒ ×ᵒ 𝒱⟦ c ⟧ V M′)) 83 | ⊎ᵒ (∃ᵒ[ N′ ] (M′ —→ N′)ᵒ ×ᵒ ▷ᵒ (ℰ⟦ c ⟧ M N′)) 84 | ⊎ᵒ (Blame M′)ᵒ) 85 | ℰ-stmt {c}{M}{M′} = 86 | let X₂ = inj₂ (c , M , M′) in 87 | ℰ⟦ c ⟧ M M′ ⩦⟨ ≡ᵒ-refl refl ⟩ 88 | μᵒ pre-ℰ⊎𝒱 X₂ ⩦⟨ fixpointᵒ pre-ℰ⊎𝒱 X₂ ⟩ 89 | # (pre-ℰ⊎𝒱 X₂) (ℰ⊎𝒱 , ttᵖ) 90 | ⩦⟨ {!!} ⟩ 91 | (((Value M′)ᵒ ×ᵒ (∃ᵒ[ V ] (M —↠ V)ᵒ ×ᵒ (Value V)ᵒ ×ᵒ 𝒱⟦ c ⟧ V M′)) 92 | ⊎ᵒ (∃ᵒ[ N′ ] (M′ —→ N′)ᵒ ×ᵒ ▷ᵒ (ℰ⟦ c ⟧ M N′)) 93 | ⊎ᵒ (Blame M′)ᵒ) 94 | ∎ 95 | {- 96 | where 97 | X₁ = λ V → (inj₁ (c , V , M′)) 98 | EQ = cong-⊎ᵒ (cong-×ᵒ (≡ᵒ-refl refl) 99 | (cong-∃ (λ V → cong-×ᵒ (≡ᵒ-refl refl) 100 | (cong-×ᵒ (≡ᵒ-refl refl) 101 | (≡ᵒ-sym (fixpointᵒ pre-ℰ⊎𝒱 (X₁ V))))))) 102 | (≡ᵒ-refl refl) 103 | -} 104 | 105 | ℰ-suc : ∀{c}{M}{M′}{k} 106 | → #(ℰ⟦ c ⟧ M M′) (suc k) ≡ 107 | #(((Value M′)ᵒ ×ᵒ (∃ᵒ[ V ] (M —↠ V)ᵒ ×ᵒ (Value V)ᵒ ×ᵒ 𝒱⟦ c ⟧ V M′)) 108 | ⊎ᵒ (∃ᵒ[ N′ ] (M′ —→ N′)ᵒ ×ᵒ ▷ᵒ (ℰ⟦ c ⟧ M N′)) 109 | ⊎ᵒ (Blame M′)ᵒ) (suc k) 110 | ℰ-suc {c}{M}{M′}{k} = refl 111 | 112 | {- Relate Open Terms -} 113 | 114 | 𝓖⟦_⟧ : (Γ : List Prec) → Subst → Subst → List Setᵒ 115 | 𝓖⟦ [] ⟧ σ σ′ = [] 116 | 𝓖⟦ c ∷ Γ ⟧ σ σ′ = (𝒱⟦ c ⟧ (σ 0) (σ′ 0)) 117 | ∷ 𝓖⟦ Γ ⟧ (λ x → σ (suc x)) (λ x → σ′ (suc x)) 118 | 119 | _⊨_⊑_⦂_ : List Prec → Term → Term → Prec → Set 120 | Γ ⊨ M ⊑ M′ ⦂ c = ∀ (γ γ′ : Subst) → 𝓖⟦ Γ ⟧ γ γ′ ⊢ᵒ ℰ⟦ c ⟧ (⟪ γ ⟫ M) (⟪ γ′ ⟫ M′) 121 | 122 | {---------- Fundamental Theorem -----------------------------------------------} 123 | 124 | fundamental : ∀ {Γ}{A}{A′}{A⊑A′ : A ⊑ A′} → (M M′ : Term) 125 | → Γ ⊩ M ⊑ M′ ⦂ A⊑A′ 126 | ---------------------------- 127 | → Γ ⊨ M ⊑ M′ ⦂ (A , A′ , A⊑A′) 128 | fundamental {Γ} {A} {A′} {A⊑A′} M⊑M′ = {!!} 129 | 130 | {---------- Gradual Guarantee -------------------------------------------------} 131 | 132 | ℰ-preserve-multi : ∀{c : Prec}{k} 133 | → (M M′ N′ : Term) 134 | → (M′→N′ : M′ —↠ N′) 135 | → #(ℰ⟦ c ⟧ M M′) (suc (len M′→N′ + k)) 136 | → #(ℰ⟦ c ⟧ M N′) (suc k) 137 | ℰ-preserve-multi{c}{k} M M′ N′ (.M′ END) ℰMM′ = ℰMM′ 138 | ℰ-preserve-multi{c}{k} M M′ N′ (.M′ —→⟨ M′→M′₁ ⟩ M′₁→N′) ℰMM′ 139 | rewrite ℰ-suc{c}{M}{M′}{suc (len M′₁→N′ + k)} 140 | with ℰMM′ 141 | ... | inj₁ (m′ , V , M→V , v , 𝒱VM′) = 142 | ⊥-elim (value-irreducible m′ M′→M′₁) 143 | ... | inj₂ (inj₂ b′) = 144 | ⊥-elim (blame-irred b′ M′→M′₁) 145 | ... | inj₂ (inj₁ (M′₂ , M′→M′₂ , ▷ℰMM′₂)) 146 | rewrite deterministic M′→M′₁ M′→M′₂ = 147 | ℰ-preserve-multi M M′₂ N′ M′₁→N′ ▷ℰMM′₂ 148 | 149 | ℰ-catchup : ∀{c}{k} 150 | → (M V′ : Term) 151 | → Value V′ 152 | → #(ℰ⟦ c ⟧ M V′) (suc k) 153 | → ∃[ V ] ((M —↠ V) × (Value V) × #(𝒱⟦ c ⟧ V V′) (suc k)) 154 | ℰ-catchup {c}{k} M V′ v′ ℰMV′ 155 | rewrite ℰ-suc{c}{M}{V′}{k} 156 | with ℰMV′ 157 | ... | inj₂ (inj₂ isBlame) = ⊥-elim (blame-not-value v′ refl) 158 | ... | inj₂ (inj₁ (V′₂ , V′→V′₂ , _)) = ⊥-elim (value-irreducible v′ V′→V′₂) 159 | ... | inj₁ (v′ , V , M→V , v , 𝒱VV′) = V , M→V , v , 𝒱VV′ 160 | 161 | {- 162 | If the more precise term goes to a value, so does the less precise term. 163 | -} 164 | GG2a : ∀{A}{A′}{A⊑A′ : A ⊑ A′}{M}{M′}{V′} 165 | → [] ⊩ M ⊑ M′ ⦂ A⊑A′ 166 | → M′ —↠ V′ 167 | → Value V′ 168 | → ∃[ V ] (M —↠ V) × (Value V) × # (𝒱⟦ A , A′ , A⊑A′ ⟧ V V′) 1 169 | GG2a {A}{A′}{A⊑A′}{M}{M′}{V′} M⊑M′ M′→V′ v′ = 170 | let ⊨M⊑M′ = fundamental M M′ M⊑M′ in 171 | let ℰMM′ = ⊢ᵒ-elim (⊨M⊑M′ id id) (suc (len M′→V′)) tt in 172 | let ℰMV′ = ℰ-preserve-multi{k = 0} M M′ V′ M′→V′ 173 | (subst (λ X → # (ℰ⟦ A , A′ , A⊑A′ ⟧ M M′) (suc X)) 174 | (sym (+-identityʳ (len M′→V′))) ℰMM′) in 175 | ℰ-catchup M V′ v′ ℰMV′ 176 | 177 | --------------------------------------------------------------------------------