3
3
4
4
import Katydid.Conal.LanguageNotation
5
5
import Mathlib.Logic.Equiv.Defs -- ≃
6
- open Lang
6
+ import Katydid.Std.Tipe2
7
+ open dLang
7
8
open List
8
9
open Char
9
10
open String
@@ -53,21 +54,21 @@ def example_of_proof_relevant_parse2 : (char 'a', (char 'b' ⋃ char 'c')) (toLi
53
54
54
55
-- ν⇃ : Lang → Set ℓ -- “nullable”
55
56
-- ν⇃ P = P []
56
- def ν (P : Lang α) : Type u := -- backslash nu
57
+ def ν' (P : dLang α) : Type u := -- backslash nu
57
58
P []
58
59
59
60
-- δ⇃ : Lang → A → Lang -- “derivative”
60
61
-- δ⇃ P a w = P (a ∷ w)
61
- def δ (P : Lang α) (a : α) : Lang α := -- backslash delta
62
+ def δ' (P : dLang α) (a : α) : dLang α := -- backslash delta
62
63
fun (w : List α) => P (a :: w)
63
64
64
- attribute [simp] ν δ
65
+ attribute [simp] ν' δ'
65
66
66
67
-- ν∅ : ν ∅ ≡ ⊥
67
68
-- ν∅ = refl
68
69
theorem nullable_emptySet :
69
70
∀ (α: Type ),
70
- @ν α ∅ ≡ PEmpty := by
71
+ @ν' α ∅ ≡ PEmpty := by
71
72
intro α
72
73
constructor
73
74
rfl
@@ -76,7 +77,7 @@ theorem nullable_emptySet:
76
77
-- ν𝒰 = refl
77
78
theorem nullable_universal :
78
79
∀ (α: Type ),
79
- @ν α 𝒰 ≡ PUnit := by
80
+ @ν' α 𝒰 ≡ PUnit := by
80
81
intro α
81
82
constructor
82
83
rfl
@@ -89,7 +90,7 @@ theorem nullable_universal:
89
90
-- (λ { refl → refl })
90
91
theorem nullable_emptyStr :
91
92
∀ (α: Type ),
92
- @ν α ε ≃ PUnit := by
93
+ @ν' α ε ≃ PUnit := by
93
94
intro α
94
95
refine Equiv.mk ?a ?b ?c ?d
95
96
intro _
@@ -105,7 +106,7 @@ theorem nullable_emptyStr:
105
106
106
107
theorem nullable_emptyStr' :
107
108
∀ (α: Type ),
108
- @ν α ε ≃ PUnit :=
109
+ @ν' α ε ≃ PUnit :=
109
110
fun _ => Equiv.mk
110
111
(fun _ => PUnit.unit)
111
112
(fun _ => by constructor; rfl)
@@ -116,7 +117,7 @@ theorem nullable_emptyStr':
116
117
-- ν` = mk↔′ (λ ()) (λ ()) (λ ()) (λ ())
117
118
theorem nullable_char :
118
119
∀ (c: α),
119
- ν (char c) ≃ PEmpty := by
120
+ ν' (char c) ≃ PEmpty := by
120
121
intro α
121
122
simp
122
123
apply Equiv.mk
@@ -131,7 +132,7 @@ theorem nullable_char:
131
132
132
133
theorem nullable_char' :
133
134
∀ (c: α),
134
- ν (char c) -> PEmpty := by
135
+ ν' (char c) -> PEmpty := by
135
136
intro
136
137
refine (fun x => ?c)
137
138
simp at x
@@ -145,26 +146,26 @@ theorem nullable_char':
145
146
-- ν∪ : ν (P ∪ Q) ≡ (ν P ⊎ ν Q)
146
147
-- ν∪ = refl
147
148
theorem nullable_or :
148
- ∀ (P Q: Lang α),
149
- ν (P ⋃ Q) ≡ (Sum (ν P) (ν Q)) := by
149
+ ∀ (P Q: dLang α),
150
+ ν' (P ⋃ Q) ≡ (Sum (ν' P) (ν' Q)) := by
150
151
intro P Q
151
152
constructor
152
153
rfl
153
154
154
155
-- ν∩ : ν (P ∩ Q) ≡ (ν P × ν Q)
155
156
-- ν∩ = refl
156
157
theorem nullable_and :
157
- ∀ (P Q: Lang α),
158
- ν (P ⋂ Q) ≡ (Prod (ν P) (ν Q)) := by
158
+ ∀ (P Q: dLang α),
159
+ ν' (P ⋂ Q) ≡ (Prod (ν' P) (ν' Q)) := by
159
160
intro P Q
160
161
constructor
161
162
rfl
162
163
163
164
-- ν· : ν (s · P) ≡ (s × ν P)
164
165
-- ν· = refl
165
166
theorem nullable_scalar :
166
- ∀ (s: Type ) (P: Lang α),
167
- ν (Lang .scalar s P) ≡ (Prod s (ν P)) := by
167
+ ∀ (s: Type ) (P: dLang α),
168
+ ν' (dLang .scalar s P) ≡ (Prod s (ν' P)) := by
168
169
intro P Q
169
170
constructor
170
171
rfl
@@ -176,8 +177,8 @@ theorem nullable_scalar:
176
177
-- (λ { (νP , νQ) → refl } )
177
178
-- (λ { (([] , []) , refl , νP , νQ) → refl})
178
179
theorem nullable_concat :
179
- ∀ (P Q: Lang α),
180
- ν (P, Q) ≃ (Prod (ν Q) (ν P)) := by
180
+ ∀ (P Q: dLang α),
181
+ ν' (P, Q) ≃ (Prod (ν' Q) (ν' P)) := by
181
182
-- TODO
182
183
sorry
183
184
@@ -210,16 +211,16 @@ theorem nullable_concat:
210
211
-- (ν P) ✶
211
212
-- ∎ where open ↔R
212
213
theorem nullable_star :
213
- ∀ (P: Lang α),
214
- ν (P *) ≃ List (ν P) := by
214
+ ∀ (P: dLang α),
215
+ ν' (P *) ≃ List (ν' P) := by
215
216
-- TODO
216
217
sorry
217
218
218
219
-- δ∅ : δ ∅ a ≡ ∅
219
220
-- δ∅ = refl
220
221
theorem derivative_emptySet :
221
222
∀ (a: α),
222
- (δ ∅ a) ≡ ∅ := by
223
+ (δ' ∅ a) ≡ ∅ := by
223
224
intro a
224
225
constructor
225
226
rfl
@@ -228,7 +229,7 @@ theorem derivative_emptySet:
228
229
-- δ𝒰 = refl
229
230
theorem derivative_universal :
230
231
∀ (a: α),
231
- (δ 𝒰 a) ≡ 𝒰 := by
232
+ (δ' 𝒰 a) ≡ 𝒰 := by
232
233
intro a
233
234
constructor
234
235
rfl
@@ -238,7 +239,7 @@ theorem derivative_universal:
238
239
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
239
240
theorem derivative_emptyStr :
240
241
∀ (a: α),
241
- (δ ε a) ≡ ∅ := by
242
+ (δ' ε a) ≡ ∅ := by
242
243
-- TODO
243
244
sorry
244
245
@@ -251,9 +252,9 @@ theorem derivative_emptyStr:
251
252
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
252
253
theorem derivative_char :
253
254
∀ (a: α) (c: α),
254
- (δ (char c) a) ≡ Lang .scalar (a ≡ c) ε := by
255
+ (δ' (char c) a) ≡ dLang .scalar (a ≡ c) ε := by
255
256
intros a c
256
- unfold δ
257
+ unfold δ'
257
258
unfold char
258
259
unfold emptyStr
259
260
unfold scalar
@@ -262,26 +263,26 @@ theorem derivative_char:
262
263
-- δ∪ : δ (P ∪ Q) a ≡ δ P a ∪ δ Q a
263
264
-- δ∪ = refl
264
265
theorem derivative_or :
265
- ∀ (a: α) (P Q: Lang α),
266
- (δ (P ⋃ Q) a) ≡ ((δ P a) ⋃ (δ Q a)) := by
266
+ ∀ (a: α) (P Q: dLang α),
267
+ (δ' (P ⋃ Q) a) ≡ ((δ' P a) ⋃ (δ' Q a)) := by
267
268
intro a P Q
268
269
constructor
269
270
rfl
270
271
271
272
-- δ∩ : δ (P ∩ Q) a ≡ δ P a ∩ δ Q a
272
273
-- δ∩ = refl
273
274
theorem derivative_and :
274
- ∀ (a: α) (P Q: Lang α),
275
- (δ (P ⋂ Q) a) ≡ ((δ P a) ⋂ (δ Q a)) := by
275
+ ∀ (a: α) (P Q: dLang α),
276
+ (δ' (P ⋂ Q) a) ≡ ((δ' P a) ⋂ (δ' Q a)) := by
276
277
intro a P Q
277
278
constructor
278
279
rfl
279
280
280
281
-- δ· : δ (s · P) a ≡ s · δ P a
281
282
-- δ· = refl
282
283
theorem derivative_scalar :
283
- ∀ (a: α) (s: Type ) (P: Lang α),
284
- (δ (Lang .scalar s P) a) ≡ (Lang .scalar s (δ P a)) := by
284
+ ∀ (a: α) (s: Type ) (P: dLang α),
285
+ (δ (dLang .scalar s P) a) ≡ (dLang .scalar s (δ' P a)) := by
285
286
intro a s P
286
287
constructor
287
288
rfl
@@ -298,9 +299,9 @@ theorem derivative_scalar:
298
299
-- ; ((.a ∷ u , v) , refl , Pu , Qv) → refl })
299
300
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
300
301
theorem derivative_concat :
301
- ∀ (a: α) (P Q: Lang α),
302
+ ∀ (a: α) (P Q: dLang α),
302
303
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
303
- (δ (P , Q) a) ≡ Lang .scalar (ν P) ((δ Q a) ⋃ ((δ P a), Q)) := by
304
+ (δ' (P , Q) a) ≡ dLang .scalar (ν' P) ((δ' Q a) ⋃ ((δ' P a), Q)) := by
304
305
-- TODO
305
306
sorry
306
307
@@ -338,8 +339,8 @@ theorem derivative_concat:
338
339
-- ∎ where open ↔R
339
340
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
340
341
theorem derivative_star :
341
- ∀ (a: α) (P: Lang α),
342
+ ∀ (a: α) (P: dLang α),
342
343
-- TODO: Redo this definition to do extensional isomorphism: `⟷` properly
343
- (δ (P *) a) ≡ Lang .scalar (List (ν P)) (δ P a, P *) := by
344
+ (δ' (P *) a) ≡ dLang .scalar (List (ν' P)) (δ' P a, P *) := by
344
345
-- TODO
345
346
sorry
0 commit comments