3 inductive re (S: DeqSet) : Type[0] ≝
7 | c: re S → re S → re S
8 | o: re S → re S → re S
11 interpretation "re epsilon" 'epsilon = (e ?).
12 interpretation "re or" 'plus a b = (o ? a b).
13 interpretation "re cat" 'middot a b = (c ? a b).
14 interpretation "re star" 'star a = (k ? a).
16 notation < "a" non associative with precedence 90 for @{ 'ps $a}.
17 notation > "` term 90 a" non associative with precedence 90 for @{ 'ps $a}.
18 interpretation "atom" 'ps a = (s ? a).
20 notation "`∅" non associative with precedence 90 for @{ 'empty }.
21 interpretation "empty" 'empty = (z ?).
23 let rec in_l (S : DeqSet) (r : re S) on r : word S → Prop ≝
28 | c r1 r2 ⇒ (in_l ? r1) · (in_l ? r2)
29 | o r1 r2 ⇒ (in_l ? r1) ∪ (in_l ? r2)
30 | k r1 ⇒ (in_l ? r1) ^*].
32 notation "\sem{term 19 E}" non associative with precedence 75 for @{'in_l $E}.
33 interpretation "in_l" 'in_l E = (in_l ? E).
34 interpretation "in_l mem" 'mem w l = (in_l ? l w).
36 lemma rsem_star : ∀S.∀r: re S. \sem{r^*} = \sem{r}^*.
41 inductive pitem (S: DeqSet) : Type[0] ≝
46 | pc: pitem S → pitem S → pitem S
47 | po: pitem S → pitem S → pitem S
48 | pk: pitem S → pitem S.
50 definition pre ≝ λS.pitem S × bool.
52 interpretation "pitem star" 'star a = (pk ? a).
53 interpretation "pitem or" 'plus a b = (po ? a b).
54 interpretation "pitem cat" 'middot a b = (pc ? a b).
55 notation < ".a" non associative with precedence 90 for @{ 'pp $a}.
56 notation > "`. term 90 a" non associative with precedence 90 for @{ 'pp $a}.
57 interpretation "pitem pp" 'pp a = (pp ? a).
58 interpretation "pitem ps" 'ps a = (ps ? a).
59 interpretation "pitem epsilon" 'epsilon = (pe ?).
60 interpretation "pitem empty" 'empty = (pz ?).
62 let rec forget (S: DeqSet) (l : pitem S) on l: re S ≝
68 | pc E1 E2 ⇒ (forget ? E1) · (forget ? E2)
69 | po E1 E2 ⇒ (forget ? E1) + (forget ? E2)
70 | pk E ⇒ (forget ? E)^* ].
72 (* notation < "|term 19 e|" non associative with precedence 70 for @{'forget $e}.*)
73 interpretation "forget" 'norm a = (forget ? a).
75 lemma erase_dot : ∀S.∀e1,e2:pitem S. |e1 · e2| = c ? (|e1|) (|e2|).
78 lemma erase_plus : ∀S.∀i1,i2:pitem S.
79 |i1 + i2| = |i1| + |i2|.
82 lemma erase_star : ∀S.∀i:pitem S.|i^*| = |i|^*.
85 (* boolean equality *)
86 let rec beqitem S (i1,i2: pitem S) on i1 ≝
88 [ pz ⇒ match i2 with [ pz ⇒ true | _ ⇒ false]
89 | pe ⇒ match i2 with [ pe ⇒ true | _ ⇒ false]
90 | ps y1 ⇒ match i2 with [ ps y2 ⇒ y1==y2 | _ ⇒ false]
91 | pp y1 ⇒ match i2 with [ pp y2 ⇒ y1==y2 | _ ⇒ false]
92 | po i11 i12 ⇒ match i2 with
93 [ po i21 i22 ⇒ beqitem S i11 i21 ∧ beqitem S i12 i22
95 | pc i11 i12 ⇒ match i2 with
96 [ pc i21 i22 ⇒ beqitem S i11 i21 ∧ beqitem S i12 i22
98 | pk i11 ⇒ match i2 with [ pk i21 ⇒ beqitem S i11 i21 | _ ⇒ false]
101 lemma beqitem_true: ∀S,i1,i2. iff (beqitem S i1 i2 = true) (i1 = i2).
103 [#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct
104 |#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct
105 |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct
106 [>(\P H) // | @(\b (refl …))]
107 |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct
108 [>(\P H) // | @(\b (refl …))]
109 |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] %
110 normalize #H destruct
111 [cases (true_or_false (beqitem S i11 i21)) #H1
112 [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H
113 |>H1 in H; normalize #abs @False_ind /2/
115 |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) //
117 |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] %
118 normalize #H destruct
119 [cases (true_or_false (beqitem S i11 i21)) #H1
120 [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H
121 |>H1 in H; normalize #abs @False_ind /2/
123 |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) //
125 |#i3 #Hind #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i4] %
126 normalize #H destruct
127 [>(proj1 … (Hind i4) H) // |>(proj2 … (Hind i4) (refl …)) //]
131 definition DeqItem ≝ λS.
132 mk_DeqSet (pitem S) (beqitem S) (beqitem_true S).
134 unification hint 0 ≔ S;
135 X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S)
136 (* ---------------------------------------- *) ⊢
139 unification hint 0 ≔ S,i1,i2;
140 X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S)
141 (* ---------------------------------------- *) ⊢
142 beqitem S i1 i2 ≡ eqb X i1 i2.
146 let rec in_pl (S : DeqSet) (r : pitem S) on r : word S → Prop ≝
152 | pc r1 r2 ⇒ (in_pl ? r1) · \sem{forget ? r2} ∪ (in_pl ? r2)
153 | po r1 r2 ⇒ (in_pl ? r1) ∪ (in_pl ? r2)
154 | pk r1 ⇒ (in_pl ? r1) · \sem{forget ? r1}^* ].
156 interpretation "in_pl" 'in_l E = (in_pl ? E).
157 interpretation "in_pl mem" 'mem w l = (in_pl ? l w).
159 definition in_prl ≝ λS : DeqSet.λp:pre S.
160 if (\snd p) then \sem{\fst p} ∪ {ϵ} else \sem{\fst p}.
162 interpretation "in_prl mem" 'mem w l = (in_prl ? l w).
163 interpretation "in_prl" 'in_l E = (in_prl ? E).
165 lemma sem_pre_true : ∀S.∀i:pitem S.
166 \sem{〈i,true〉} = \sem{i} ∪ {ϵ}.
169 lemma sem_pre_false : ∀S.∀i:pitem S.
170 \sem{〈i,false〉} = \sem{i}.
173 lemma sem_cat: ∀S.∀i1,i2:pitem S.
174 \sem{i1 · i2} = \sem{i1} · \sem{|i2|} ∪ \sem{i2}.
177 lemma sem_cat_w: ∀S.∀i1,i2:pitem S.∀w.
178 \sem{i1 · i2} w = ((\sem{i1} · \sem{|i2|}) w ∨ \sem{i2} w).
181 lemma sem_plus: ∀S.∀i1,i2:pitem S.
182 \sem{i1 + i2} = \sem{i1} ∪ \sem{i2}.
185 lemma sem_plus_w: ∀S.∀i1,i2:pitem S.∀w.
186 \sem{i1 + i2} w = (\sem{i1} w ∨ \sem{i2} w).
189 lemma sem_star : ∀S.∀i:pitem S.
190 \sem{i^*} = \sem{i} · \sem{|i|}^*.
193 lemma sem_star_w : ∀S.∀i:pitem S.∀w.
194 \sem{i^*} w = (∃w1,w2.w1 @ w2 = w ∧ \sem{i} w1 ∧ \sem{|i|}^* w2).
197 lemma append_eq_nil : ∀S.∀w1,w2:word S. w1 @ w2 = ϵ → w1 = ϵ.
198 #S #w1 #w2 cases w1 // #a #tl normalize #H destruct qed.
200 lemma not_epsilon_lp : ∀S:DeqSet.∀e:pitem S. ¬ (ϵ ∈ e).
201 #S #e elim e normalize /2/
202 [#r1 #r2 * #n1 #n2 % * /2/ * #w1 * #w2 * * #H
203 >(append_eq_nil …H…) /2/
204 |#r1 #r2 #n1 #n2 % * /2/
205 |#r #n % * #w1 * #w2 * * #H >(append_eq_nil …H…) /2/
210 lemma epsilon_to_true : ∀S.∀e:pre S. ϵ ∈ e → \snd e = true.
211 #S * #i #b cases b // normalize #H @False_ind /2/
214 lemma true_to_epsilon : ∀S.∀e:pre S. \snd e = true → ϵ ∈ e.
215 #S * #i #b #btrue normalize in btrue; >btrue %2 //
218 lemma minus_eps_item: ∀S.∀i:pitem S. \sem{i} =1 \sem{i}-{[ ]}.
220 [#H whd % // normalize @(not_to_not … (not_epsilon_lp …i)) //
225 lemma minus_eps_pre: ∀S.∀e:pre S. \sem{\fst e} =1 \sem{e}-{[ ]}.
227 [>sem_pre_true normalize in ⊢ (??%?); #w %
228 [/3/ | * * // #H1 #H2 @False_ind @(absurd …H1 H2)]
229 |>sem_pre_false normalize in ⊢ (??%?); #w % [ /3/ | * // ]
233 definition lo ≝ λS:DeqSet.λa,b:pre S.〈\fst a + \fst b,\snd a ∨ \snd b〉.
234 notation "a ⊕ b" left associative with precedence 65 for @{'oplus $a $b}.
235 interpretation "oplus" 'oplus a b = (lo ? a b).
237 lemma lo_def: ∀S.∀i1,i2:pitem S.∀b1,b2. 〈i1,b1〉⊕〈i2,b2〉=〈i1+i2,b1∨b2〉.
240 definition pre_concat_r ≝ λS:DeqSet.λi:pitem S.λe:pre S.
241 match e with [ mk_Prod i1 b ⇒ 〈i · i1, b〉].
243 notation "i ◃ e" left associative with precedence 65 for @{'lhd $i $e}.
244 interpretation "pre_concat_r" 'lhd i e = (pre_concat_r ? i e).
246 lemma eq_to_ex_eq: ∀S.∀A,B:word S → Prop.
248 #S #A #B #H >H /2/ qed.
250 lemma sem_pre_concat_r : ∀S,i.∀e:pre S.
251 \sem{i ◃ e} =1 \sem{i} · \sem{|\fst e|} ∪ \sem{e}.
252 #S #i * #i1 #b1 cases b1 [2: @eq_to_ex_eq //]
253 >sem_pre_true >sem_cat >sem_pre_true /2/
256 definition pre_concat_l ≝ λS:DeqSet.λbcast:∀S:DeqSet.pitem S → pre S.λe1:pre S.λi2:pitem S.
258 [ mk_Prod i1 b1 ⇒ match b1 with
259 [ true ⇒ (i1 ◃ (bcast ? i2))
260 | false ⇒ 〈i1 · i2,false〉
264 notation "a ▹ b" left associative with precedence 65 for @{'tril eclose $a $b}.
265 interpretation "item-pre concat" 'tril op a b = (pre_concat_l ? op a b).
267 notation "•" non associative with precedence 65 for @{eclose ?}.
269 let rec eclose (S: DeqSet) (i: pitem S) on i : pre S ≝
273 | ps x ⇒ 〈 `.x, false〉
274 | pp x ⇒ 〈 `.x, false 〉
275 | po i1 i2 ⇒ •i1 ⊕ •i2
276 | pc i1 i2 ⇒ •i1 ▹ i2
277 | pk i ⇒ 〈(\fst (•i))^*,true〉].
279 notation "• x" non associative with precedence 65 for @{'eclose $x}.
280 interpretation "eclose" 'eclose x = (eclose ? x).
282 lemma eclose_plus: ∀S:DeqSet.∀i1,i2:pitem S.
283 •(i1 + i2) = •i1 ⊕ •i2.
286 lemma eclose_dot: ∀S:DeqSet.∀i1,i2:pitem S.
287 •(i1 · i2) = •i1 ▹ i2.
290 lemma eclose_star: ∀S:DeqSet.∀i:pitem S.
291 •i^* = 〈(\fst(•i))^*,true〉.
294 definition lift ≝ λS.λf:pitem S →pre S.λe:pre S.
296 [ mk_Prod i b ⇒ 〈\fst (f i), \snd (f i) ∨ b〉].
298 definition preclose ≝ λS. lift S (eclose S).
299 interpretation "preclose" 'eclose x = (preclose ? x).
302 lemma sem_oplus: ∀S:DeqSet.∀e1,e2:pre S.
303 \sem{e1 ⊕ e2} =1 \sem{e1} ∪ \sem{e2}.
304 #S * #i1 #b1 * #i2 #b2 #w %
305 [cases b1 cases b2 normalize /2/ * /3/ * /3/
306 |cases b1 cases b2 normalize /2/ * /3/ * /3/
312 〈i1,true〉 ▹ i2 = i1 ◃ (•i2).
315 lemma odot_true_bis :
317 〈i1,true〉 ▹ i2 = 〈i1 · \fst (•i2), \snd (•i2)〉.
318 #S #i1 #i2 normalize cases (•i2) // qed.
322 〈i1,false〉 ▹ i2 = 〈i1 · i2, false〉.
325 lemma LcatE : ∀S.∀e1,e2:pitem S.
326 \sem{e1 · e2} = \sem{e1} · \sem{|e2|} ∪ \sem{e2}.
329 lemma erase_bull : ∀S.∀i:pitem S. |\fst (•i)| = |i|.
331 [ #i1 #i2 #IH1 #IH2 >erase_dot <IH1 >eclose_dot
332 cases (•i1) #i11 #b1 cases b1 // <IH2 >odot_true_bis //
333 | #i1 #i2 #IH1 #IH2 >eclose_plus >(erase_plus … i1) <IH1 <IH2
334 cases (•i1) #i11 #b1 cases (•i2) #i21 #b2 //
335 | #i #IH >eclose_star >(erase_star … i) <IH cases (•i) //
340 lemma sem_eclose_star: ∀S:DeqSet.∀i:pitem S.
341 \sem{〈i^*,true〉} =1 \sem{〈i,false〉}·\sem{|i|}^* ∪ {ϵ}.
345 (* theorem 16: 1 → 3 *)
346 lemma odot_dot_aux : ∀S.∀e1:pre S.∀i2:pitem S.
347 \sem{•i2} =1 \sem{i2} ∪ \sem{|i2|} →
348 \sem{e1 ▹ i2} =1 \sem{e1} · \sem{|i2|} ∪ \sem{i2}.
349 #S * #i1 #b1 #i2 cases b1
350 [2:#th >odot_false >sem_pre_false >sem_pre_false >sem_cat /2/
351 |#H >odot_true >sem_pre_true @(eqP_trans … (sem_pre_concat_r …))
352 >erase_bull @eqP_trans [|@(eqP_union_l … H)]
353 @eqP_trans [|@eqP_union_l[|@union_comm ]]
354 @eqP_trans [|@eqP_sym @union_assoc ] /3/
358 lemma minus_eps_pre_aux: ∀S.∀e:pre S.∀i:pitem S.∀A.
359 \sem{e} =1 \sem{i} ∪ A → \sem{\fst e} =1 \sem{i} ∪ (A - {[ ]}).
361 @eqP_trans [|@minus_eps_pre]
362 @eqP_trans [||@eqP_union_r [|@eqP_sym @minus_eps_item]]
363 @eqP_trans [||@distribute_substract]
368 theorem sem_bull: ∀S:DeqSet. ∀i:pitem S. \sem{•i} =1 \sem{i} ∪ \sem{|i|}.
370 [#w normalize % [/2/ | * //]
372 |#x normalize #w % [ /2/ | * [@False_ind | //]]
373 |#x normalize #w % [ /2/ | * // ]
374 |#i1 #i2 #IH1 #IH2 >eclose_dot
375 @eqP_trans [|@odot_dot_aux //] >sem_cat
378 [|@eqP_trans [|@(cat_ext_l … IH1)] @distr_cat_r]]
379 @eqP_trans [|@union_assoc]
380 @eqP_trans [||@eqP_sym @union_assoc]
382 |#i1 #i2 #IH1 #IH2 >eclose_plus
383 @eqP_trans [|@sem_oplus] >sem_plus >erase_plus
384 @eqP_trans [|@(eqP_union_l … IH2)]
385 @eqP_trans [|@eqP_sym @union_assoc]
386 @eqP_trans [||@union_assoc] @eqP_union_r
387 @eqP_trans [||@eqP_sym @union_assoc]
388 @eqP_trans [||@eqP_union_l [|@union_comm]]
389 @eqP_trans [||@union_assoc] /2/
390 |#i #H >sem_pre_true >sem_star >erase_bull >sem_star
391 @eqP_trans [|@eqP_union_r [|@cat_ext_l [|@minus_eps_pre_aux //]]]
392 @eqP_trans [|@eqP_union_r [|@distr_cat_r]]
393 @eqP_trans [|@union_assoc] @eqP_union_l >erase_star
394 @eqP_sym @star_fix_eps
399 let rec blank (S: DeqSet) (i: re S) on i :pitem S ≝
404 | o e1 e2 ⇒ (blank S e1) + (blank S e2)
405 | c e1 e2 ⇒ (blank S e1) · (blank S e2)
406 | k e ⇒ (blank S e)^* ].
408 lemma forget_blank: ∀S.∀e:re S.|blank S e| = e.
409 #S #e elim e normalize //
412 lemma sem_blank: ∀S.∀e:re S.\sem{blank S e} =1 ∅.
416 |#e1 #e2 #Hind1 #Hind2 >sem_cat
417 @eqP_trans [||@(union_empty_r … ∅)]
418 @eqP_trans [|@eqP_union_l[|@Hind2]] @eqP_union_r
419 @eqP_trans [||@(cat_empty_l … ?)] @cat_ext_l @Hind1
420 |#e1 #e2 #Hind1 #Hind2 >sem_plus
421 @eqP_trans [||@(union_empty_r … ∅)]
422 @eqP_trans [|@eqP_union_l[|@Hind2]] @eqP_union_r @Hind1
424 @eqP_trans [||@(cat_empty_l … ?)] @cat_ext_l @Hind
428 theorem re_embedding: ∀S.∀e:re S.
429 \sem{•(blank S e)} =1 \sem{e}.
430 #S #e @eqP_trans [|@sem_bull] >forget_blank
431 @eqP_trans [|@eqP_union_r [|@sem_blank]]
432 @eqP_trans [|@union_comm] @union_empty_r.
435 (* lefted operations *)
436 definition lifted_cat ≝ λS:DeqSet.λe:pre S.
437 lift S (pre_concat_l S eclose e).
439 notation "e1 ⊙ e2" left associative with precedence 70 for @{'odot $e1 $e2}.
441 interpretation "lifted cat" 'odot e1 e2 = (lifted_cat ? e1 e2).
443 lemma odot_true_b : ∀S.∀i1,i2:pitem S.∀b.
444 〈i1,true〉 ⊙ 〈i2,b〉 = 〈i1 · (\fst (•i2)),\snd (•i2) ∨ b〉.
445 #S #i1 #i2 #b normalize in ⊢ (??%?); cases (•i2) //
448 lemma odot_false_b : ∀S.∀i1,i2:pitem S.∀b.
449 〈i1,false〉 ⊙ 〈i2,b〉 = 〈i1 · i2 ,b〉.
453 lemma erase_odot:∀S.∀e1,e2:pre S.
454 |\fst (e1 ⊙ e2)| = |\fst e1| · (|\fst e2|).
455 #S * #i1 * * #i2 #b2 // >odot_true_b >erase_dot //
458 definition lk ≝ λS:DeqSet.λe:pre S.
462 [true ⇒ 〈(\fst (eclose ? i1))^*, true〉
463 |false ⇒ 〈i1^*,false〉
467 (* notation < "a \sup ⊛" non associative with precedence 90 for @{'lk $a}.*)
468 interpretation "lk" 'lk a = (lk ? a).
469 notation "a^⊛" non associative with precedence 90 for @{'lk $a}.
472 lemma ostar_true: ∀S.∀i:pitem S.
473 〈i,true〉^⊛ = 〈(\fst (•i))^*, true〉.
476 lemma ostar_false: ∀S.∀i:pitem S.
477 〈i,false〉^⊛ = 〈i^*, false〉.
480 lemma erase_ostar: ∀S.∀e:pre S.
481 |\fst (e^⊛)| = |\fst e|^*.
484 lemma sem_odot_true: ∀S:DeqSet.∀e1:pre S.∀i.
485 \sem{e1 ⊙ 〈i,true〉} =1 \sem{e1 ▹ i} ∪ { [ ] }.
487 cut (e1 ⊙ 〈i,true〉 = 〈\fst (e1 ▹ i), \snd(e1 ▹ i) ∨ true〉) [//]
488 #H >H cases (e1 ▹ i) #i1 #b1 cases b1
489 [>sem_pre_true @eqP_trans [||@eqP_sym @union_assoc]
495 lemma eq_odot_false: ∀S:DeqSet.∀e1:pre S.∀i.
496 e1 ⊙ 〈i,false〉 = e1 ▹ i.
498 cut (e1 ⊙ 〈i,false〉 = 〈\fst (e1 ▹ i), \snd(e1 ▹ i) ∨ false〉) [//]
499 cases (e1 ▹ i) #i1 #b1 cases b1 #H @H
503 ∀S.∀e1,e2: pre S. \sem{e1 ⊙ e2} =1 \sem{e1}· \sem{|\fst e2|} ∪ \sem{e2}.
506 @eqP_trans [|@sem_odot_true]
507 @eqP_trans [||@union_assoc] @eqP_union_r @odot_dot_aux //
508 |>sem_pre_false >eq_odot_false @odot_dot_aux //
513 theorem sem_ostar: ∀S.∀e:pre S.
514 \sem{e^⊛} =1 \sem{e} · \sem{|\fst e|}^*.
516 [>sem_pre_true >sem_pre_true >sem_star >erase_bull
517 @eqP_trans [|@eqP_union_r[|@cat_ext_l [|@minus_eps_pre_aux //]]]
518 @eqP_trans [|@eqP_union_r [|@distr_cat_r]]
519 @eqP_trans [||@eqP_sym @distr_cat_r]
520 @eqP_trans [|@union_assoc] @eqP_union_l
521 @eqP_trans [||@eqP_sym @epsilon_cat_l] @eqP_sym @star_fix_eps
522 |>sem_pre_false >sem_pre_false >sem_star /2/