Skip to content

Commit 53f4211

Browse files
nadMatthewDaggitt
authored andcommitted
Add some lemmas related to renamings and substitutions (#1750)
1 parent 8fd57cc commit 53f4211

File tree

6 files changed

+294
-86
lines changed

6 files changed

+294
-86
lines changed

CHANGELOG.md

Lines changed: 25 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1117,6 +1117,10 @@ Deprecated modules
11171117

11181118
* The module `Data.Nat.Properties.Core` has been deprecated, and its one entry moved to `Data.Nat.Properties`
11191119

1120+
### Deprecation of `Data.Fin.Substitution.Example`
1121+
1122+
* The module `Data.Fin.Substitution.Example` has been deprecated, and moved to `README.Data.Fin.Substitution.UntypedLambda`
1123+
11201124
### Deprecation of `Data.Product.Function.Dependent.Setoid.WithK`
11211125

11221126
* This module has been deprecated, as none of its contents actually depended on axiom K. The contents has been moved to `Data.Product.Function.Dependent.Setoid`.
@@ -2201,6 +2205,27 @@ Additions to existing modules
22012205
inject≤-irrelevant : inject≤ i m≤n ≡ inject≤ i m≤n′
22022206
```
22032207

2208+
* Changed the fixity of `Data.Fin.Substitution.TermSubst._/Var_`.
2209+
```agda
2210+
infix 8 ↦ infixl 8
2211+
```
2212+
2213+
* Added new lemmas in `Data.Fin.Substitution.Lemmas.TermLemmas`:
2214+
```
2215+
map-var≡ : {ρ₁ : Sub Fin m n} {ρ₂ : Sub T m n} {f : Fin m → Fin n} →
2216+
(∀ x → lookup ρ₁ x ≡ f x) →
2217+
(∀ x → lookup ρ₂ x ≡ T.var (f x)) →
2218+
map T.var ρ₁ ≡ ρ₂
2219+
wk≡wk : map T.var VarSubst.wk ≡ T.wk {n = n}
2220+
id≡id : map T.var VarSubst.id ≡ T.id {n = n}
2221+
sub≡sub : {x : Fin n} → map T.var (VarSubst.sub x) ≡ T.sub (T.var x)
2222+
↑≡↑ : {ρ : Sub Fin m n} → map T.var (ρ VarSubst.↑) ≡ map T.var ρ T.↑
2223+
/Var≡/ : {ρ : Sub Fin m n} {t} → t /Var ρ ≡ t T./ map T.var ρ
2224+
sub-renaming-commutes : {ρ : Sub T m n} →
2225+
t /Var VarSubst.sub x T./ ρ ≡ t T./ ρ T.↑ T./ T.sub (lookup ρ x)
2226+
sub-commutes-with-renaming : {ρ : Sub Fin m n} →
2227+
t T./ T.sub t′ /Var ρ ≡ t /Var ρ VarSubst.↑ T./ T.sub (t′ /Var ρ)
2228+
22042229
* Added new functions in `Data.Integer.Base`:
22052230
```
22062231
_^_ : ℤ → ℕ → ℤ
Lines changed: 103 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,103 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- An example of how Data.Fin.Substitution can be used: a definition
5+
-- of substitution for the untyped λ-calculus, along with some lemmas
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --cubical-compatible --safe #-}
9+
10+
module README.Data.Fin.Substitution.UntypedLambda where
11+
12+
open import Data.Fin.Substitution
13+
open import Data.Fin.Substitution.Lemmas
14+
open import Data.Nat.Base hiding (_/_)
15+
open import Data.Fin.Base using (Fin)
16+
open import Data.Vec.Base
17+
open import Relation.Binary.PropositionalEquality.Core
18+
using (_≡_; refl; sym; cong; cong₂; module ≡-Reasoning)
19+
open import Relation.Binary.Construct.Closure.ReflexiveTransitive
20+
using (Star; ε; _◅_)
21+
22+
open ≡-Reasoning
23+
24+
private
25+
variable
26+
m n :
27+
28+
------------------------------------------------------------------------
29+
-- A representation of the untyped λ-calculus. Uses de Bruijn indices.
30+
31+
infixl 9 _·_
32+
33+
data Lam (n : ℕ) : Set where
34+
var : (x : Fin n) Lam n
35+
ƛ : (t : Lam (suc n)) Lam n
36+
_·_ : (t₁ t₂ : Lam n) Lam n
37+
38+
------------------------------------------------------------------------
39+
-- Code for applying substitutions.
40+
41+
module LamApp {ℓ} {T : Set ℓ} (l : Lift T Lam) where
42+
open Lift l hiding (var)
43+
44+
-- Applies a substitution to a term.
45+
46+
infixl 8 _/_
47+
48+
_/_ : Lam m Sub T m n Lam n
49+
var x / ρ = lift (lookup ρ x)
50+
ƛ t / ρ = ƛ (t / ρ ↑)
51+
t₁ · t₂ / ρ = (t₁ / ρ) · (t₂ / ρ)
52+
53+
open Application (record { _/_ = _/_ }) using (_/✶_)
54+
55+
-- Some lemmas about _/_.
56+
57+
ƛ-/✶-↑✶ : k {t} (ρs : Subs T m n)
58+
ƛ t /✶ ρs ↑✶ k ≡ ƛ (t /✶ ρs ↑✶ suc k)
59+
ƛ-/✶-↑✶ k ε = refl
60+
ƛ-/✶-↑✶ k (ρ ◅ ρs) = cong (_/ _) (ƛ-/✶-↑✶ k ρs)
61+
62+
·-/✶-↑✶ : k {t₁ t₂} (ρs : Subs T m n)
63+
t₁ · t₂ /✶ ρs ↑✶ k ≡ (t₁ /✶ ρs ↑✶ k) · (t₂ /✶ ρs ↑✶ k)
64+
·-/✶-↑✶ k ε = refl
65+
·-/✶-↑✶ k (ρ ◅ ρs) = cong (_/ _) (·-/✶-↑✶ k ρs)
66+
67+
lamSubst : TermSubst Lam
68+
lamSubst = record { var = var; app = LamApp._/_ }
69+
70+
open TermSubst lamSubst hiding (var)
71+
72+
------------------------------------------------------------------------
73+
-- Substitution lemmas.
74+
75+
lamLemmas : TermLemmas Lam
76+
lamLemmas = record
77+
{ termSubst = lamSubst
78+
; app-var = refl
79+
; /✶-↑✶ = Lemma./✶-↑✶
80+
}
81+
where
82+
module Lemma {T₁ T₂} {lift₁ : Lift T₁ Lam} {lift₂ : Lift T₂ Lam} where
83+
84+
open Lifted lift₁ using () renaming (_↑✶_ to _↑✶₁_; _/✶_ to _/✶₁_)
85+
open Lifted lift₂ using () renaming (_↑✶_ to _↑✶₂_; _/✶_ to _/✶₂_)
86+
87+
/✶-↑✶ : (ρs₁ : Subs T₁ m n) (ρs₂ : Subs T₂ m n)
88+
( k x var x /✶₁ ρs₁ ↑✶₁ k ≡ var x /✶₂ ρs₂ ↑✶₂ k)
89+
k t t /✶₁ ρs₁ ↑✶₁ k ≡ t /✶₂ ρs₂ ↑✶₂ k
90+
/✶-↑✶ ρs₁ ρs₂ hyp k (var x) = hyp k x
91+
/✶-↑✶ ρs₁ ρs₂ hyp k (ƛ t) = begin
92+
ƛ t /✶₁ ρs₁ ↑✶₁ k ≡⟨ LamApp.ƛ-/✶-↑✶ _ k ρs₁ ⟩
93+
ƛ (t /✶₁ ρs₁ ↑✶₁ suc k) ≡⟨ cong ƛ (/✶-↑✶ ρs₁ ρs₂ hyp (suc k) t) ⟩
94+
ƛ (t /✶₂ ρs₂ ↑✶₂ suc k) ≡⟨ sym (LamApp.ƛ-/✶-↑✶ _ k ρs₂) ⟩
95+
ƛ t /✶₂ ρs₂ ↑✶₂ k ∎
96+
/✶-↑✶ ρs₁ ρs₂ hyp k (t₁ · t₂) = begin
97+
t₁ · t₂ /✶₁ ρs₁ ↑✶₁ k ≡⟨ LamApp.·-/✶-↑✶ _ k ρs₁ ⟩
98+
(t₁ /✶₁ ρs₁ ↑✶₁ k) · (t₂ /✶₁ ρs₁ ↑✶₁ k) ≡⟨ cong₂ _·_ (/✶-↑✶ ρs₁ ρs₂ hyp k t₁)
99+
(/✶-↑✶ ρs₁ ρs₂ hyp k t₂) ⟩
100+
(t₁ /✶₂ ρs₂ ↑✶₂ k) · (t₂ /✶₂ ρs₂ ↑✶₂ k) ≡⟨ sym (LamApp.·-/✶-↑✶ _ k ρs₂) ⟩
101+
t₁ · t₂ /✶₂ ρs₂ ↑✶₂ k ∎
102+
103+
open TermLemmas lamLemmas public hiding (var)

src/Data/Fin/Substitution.agda

Lines changed: 30 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -4,12 +4,12 @@
44
-- Substitutions
55
------------------------------------------------------------------------
66

7-
-- Uses a variant of Conor McBride's technique from "Type-Preserving
8-
-- Renaming and Substitution".
7+
-- Uses a variant of Conor McBride's technique from his paper
8+
-- "Type-Preserving Renaming and Substitution".
99

10-
-- See Data.Fin.Substitution.Example for an example of how this module
11-
-- can be used: a definition of substitution for the untyped
12-
-- λ-calculus.
10+
-- See README.Data.Fin.Substitution.UntypedLambda for an example
11+
-- of how this module can be used: a definition of substitution for
12+
-- the untyped λ-calculus.
1313

1414
{-# OPTIONS --cubical-compatible --safe #-}
1515

@@ -24,23 +24,29 @@ open import Relation.Binary.Construct.Closure.ReflexiveTransitive
2424
open import Level using (Level; _⊔_; 0ℓ)
2525
open import Relation.Unary using (Pred)
2626

27+
private
28+
variable
29+
ℓ ℓ₁ ℓ₂ : Level
30+
k m n o :
31+
32+
2733
------------------------------------------------------------------------
2834
-- General functionality
2935

3036
-- A Sub T m n is a substitution which, when applied to something with
3137
-- at most m variables, yields something with at most n variables.
3238

33-
Sub : {ℓ} Pred ℕ ℓ Set
39+
Sub : Pred ℕ ℓ Set
3440
Sub T m n = Vec (T n) m
3541

3642
-- A /reversed/ sequence of matching substitutions.
3743

38-
Subs : {ℓ} Pred ℕ ℓ Set
44+
Subs : Pred ℕ ℓ Set
3945
Subs T = flip (Star (flip (Sub T)))
4046

4147
-- Some simple substitutions.
4248

43-
record Simple {ℓ : Level} (T : Pred ℕ ℓ) : Setwhere
49+
record Simple (T : Pred ℕ ℓ) : Setwhere
4450
infix 10 _↑
4551
infixl 10 _↑⋆_ _↑✶_
4652

@@ -50,40 +56,39 @@ record Simple {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
5056

5157
-- Lifting.
5258

53-
_↑ : {m n} Sub T m n Sub T (suc m) (suc n)
59+
_↑ : Sub T m n Sub T (suc m) (suc n)
5460
ρ ↑ = var zero ∷ map weaken ρ
5561

56-
_↑⋆_ : {m n} Sub T m n k Sub T (k + m) (k + n)
62+
_↑⋆_ : Sub T m n k Sub T (k + m) (k + n)
5763
ρ ↑⋆ zero = ρ
5864
ρ ↑⋆ suc k = (ρ ↑⋆ k) ↑
5965

60-
_↑✶_ : {m n} Subs T m n k Subs T (k + m) (k + n)
66+
_↑✶_ : Subs T m n k Subs T (k + m) (k + n)
6167
ρs ↑✶ k = Star.gmap (_+_ k) (λ ρ ρ ↑⋆ k) ρs
6268

6369
-- The identity substitution.
6470

65-
id : {n} Sub T n n
71+
id : Sub T n n
6672
id {zero} = []
6773
id {suc n} = id ↑
6874

6975
-- Weakening.
7076

71-
wk⋆ : k {n} Sub T n (k + n)
77+
wk⋆ : k Sub T n (k + n)
7278
wk⋆ zero = id
7379
wk⋆ (suc k) = map weaken (wk⋆ k)
7480

75-
wk : {n} Sub T n (suc n)
81+
wk : Sub T n (suc n)
7682
wk = wk⋆ 1
7783

7884
-- A substitution which only replaces the first variable.
7985

80-
sub : {n} T n Sub T (suc n) n
86+
sub : T n Sub T (suc n) n
8187
sub t = t ∷ id
8288

8389
-- Application of substitutions.
8490

85-
record Application {ℓ₁ ℓ₂ : Level} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) :
86-
Set (ℓ₁ ⊔ ℓ₂) where
91+
record Application (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) : Set (ℓ₁ ⊔ ℓ₂) where
8792
infixl 8 _/_ _/✶_
8893
infixl 9 _⊙_
8994

@@ -93,17 +98,17 @@ record Application {ℓ₁ ℓ₂ : Level} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred
9398

9499
-- Reverse composition. (Fits well with post-application.)
95100

96-
_⊙_ : {m n k} Sub T₁ m n Sub T₂ n k Sub T₁ m k
97-
ρ₁ ⊙ ρ₂ = map (λ t t / ρ₂) ρ₁
101+
_⊙_ : Sub T₁ m n Sub T₂ n o Sub T₁ m o
102+
ρ₁ ⊙ ρ₂ = map (_/ ρ₂) ρ₁
98103

99104
-- Application of multiple substitutions.
100105

101-
_/✶_ : {m n} T₁ m Subs T₂ m n T₁ n
106+
_/✶_ : T₁ m Subs T₂ m n T₁ n
102107
_/✶_ = Star.gfold Fun.id _ (flip _/_) {k = zero}
103108

104109
-- A combination of the two records above.
105110

106-
record Subst {ℓ : Level} (T : Pred ℕ ℓ) : Setwhere
111+
record Subst (T : Pred ℕ ℓ) : Setwhere
107112
field
108113
simple : Simple T
109114
application : Application T T
@@ -113,7 +118,7 @@ record Subst {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
113118

114119
-- Composition of multiple substitutions.
115120

116-
: {m n} Subs T m n Sub T m n
121+
: Subs T m n Sub T m n
117122
⨀ ε = id
118123
⨀ (ρ ◅ ε) = ρ -- Convenient optimisation; simplifies some proofs.
119124
⨀ (ρ ◅ ρs) = ⨀ ρs ⊙ ρ
@@ -123,8 +128,7 @@ record Subst {ℓ : Level} (T : Pred ℕ ℓ) : Set ℓ where
123128

124129
-- Liftings from T₁ to T₂.
125130

126-
record Lift {ℓ₁ ℓ₂ : Level} (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) :
127-
Set (ℓ₁ ⊔ ℓ₂) where
131+
record Lift (T₁ : Pred ℕ ℓ₁) (T₂ : Pred ℕ ℓ₂) : Set (ℓ₁ ⊔ ℓ₂) where
128132
field
129133
simple : Simple T₁
130134
lift : {n} T₁ n T₂ n
@@ -160,9 +164,9 @@ record TermSubst (T : Pred ℕ 0ℓ) : Set₁ where
160164
varLift : Lift Fin T
161165
varLift = record { simple = VarSubst.simple; lift = var }
162166

163-
infix 8 _/Var_
167+
infixl 8 _/Var_
164168

165-
_/Var_ : {m n} T m Sub Fin m n T n
169+
_/Var_ : T m Sub Fin m n T n
166170
_/Var_ = app varLift
167171

168172
simple : Simple T

src/Data/Fin/Substitution/Example.agda

Lines changed: 8 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,14 +1,20 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- An example of how Data.Fin.Substitution can be used: a definition
5-
-- of substitution for the untyped λ-calculus, along with some lemmas
4+
-- This module is DEPRECATED. Please see
5+
-- `README.Data.Nat.Fin.Substitution.UntypedLambda` instead.
6+
--
67
------------------------------------------------------------------------
78

89
{-# OPTIONS --cubical-compatible --safe #-}
910

1011
module Data.Fin.Substitution.Example where
1112

13+
{-# WARNING_ON_IMPORT
14+
"Data.Fin.Substitution.Example was deprecated in v2.0.
15+
Please see README.Data.Nat.Fin.Substitution.UntypedLambda instead."
16+
#-}
17+
1218
open import Data.Fin.Substitution
1319
open import Data.Fin.Substitution.Lemmas
1420
open import Data.Nat.Base hiding (_/_)

0 commit comments

Comments
 (0)