1 include "basics/list.ma".
2 include "basics/sets.ma".
3 include "basics/deqsets.ma".
5 definition word ≝ λS:DeqSet.list S.
7 notation "ϵ" non associative with precedence 90 for @{ 'epsilon }.
8 interpretation "epsilon" 'epsilon = (nil ?).
11 definition cat : ∀S,l1,l2,w.Prop ≝
12 λS.λl1,l2.λw:word S.∃w1,w2.w1 @ w2 = w ∧ l1 w1 ∧ l2 w2.
13 notation "a · b" non associative with precedence 65 for @{ 'middot $a $b}.
14 interpretation "cat lang" 'middot a b = (cat ? a b).
16 let rec flatten (S : DeqSet) (l : list (word S)) on l : word S ≝
17 match l with [ nil ⇒ [ ] | cons w tl ⇒ w @ flatten ? tl ].
19 let rec conjunct (S : DeqSet) (l : list (word S)) (r : word S → Prop) on l: Prop ≝
20 match l with [ nil ⇒ True | cons w tl ⇒ r w ∧ conjunct ? tl r ].
23 definition star ≝ λS.λl.λw:word S.∃lw.flatten ? lw = w ∧ conjunct ? lw l.
24 notation "a ^ *" non associative with precedence 90 for @{ 'star $a}.
25 interpretation "star lang" 'star l = (star ? l).
27 lemma cat_ext_l: ∀S.∀A,B,C:word S →Prop.
28 A =1 C → A · B =1 C · B.
29 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
33 lemma cat_ext_r: ∀S.∀A,B,C:word S →Prop.
34 B =1 C → A · B =1 A · C.
35 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
39 lemma cat_empty_l: ∀S.∀A:word S→Prop. ∅ · A =1 ∅.
40 #S #A #w % [|*] * #w1 * #w2 * * #_ *
43 lemma distr_cat_r: ∀S.∀A,B,C:word S →Prop.
44 (A ∪ B) · C =1 A · C ∪ B · C.
46 [* #w1 * #w2 * * #eqw * /6/ |* * #w1 * #w2 * * /6/]
51 definition deriv ≝ λS.λA:word S → Prop.λa,w. A (a::w).
53 lemma deriv_middot: ∀S,A,B,a. ¬ A ϵ →
54 deriv S (A·B) a =1 (deriv S A a) · B.
55 #S #A #B #a #noteps #w normalize %
57 [* #w2 * * #_ #Aeps @False_ind /2/
58 |#b #w2 * #w3 * * whd in ⊢ ((??%?)→?); #H destruct
59 #H #H1 @(ex_intro … w2) @(ex_intro … w3) % // % //
61 |* #w1 * #w2 * * #H #H1 #H2 @(ex_intro … (a::w1))
62 @(ex_intro … w2) % // % normalize //
67 lemma espilon_in_star: ∀S.∀A:word S → Prop.
69 #S #A @(ex_intro … [ ]) normalize /2/
72 lemma cat_to_star:∀S.∀A:word S → Prop.
73 ∀w1,w2. A w1 → A^* w2 → A^* (w1@w2).
74 #S #A #w1 #w2 #Aw * #l * #H #H1 @(ex_intro … (w1::l))
78 lemma fix_star: ∀S.∀A:word S → Prop.
81 [* #l generalize in match w; -w cases l [normalize #w * /2/]
82 #w1 #tl #w * whd in ⊢ ((??%?)→?); #eqw whd in ⊢ (%→?); *
83 #w1A #cw1 %1 @(ex_intro … w1) @(ex_intro … (flatten S tl))
84 % /2/ whd @(ex_intro … tl) /2/
85 |* [2: whd in ⊢ (%→?); #eqw <eqw //]
86 * #w1 * #w2 * * #eqw <eqw @cat_to_star
90 lemma star_fix_eps : ∀S.∀A:word S → Prop.
91 A^* =1 (A - {ϵ}) · A^* ∪ {ϵ}.
94 [* whd in ⊢ ((??%?)→?); #eqw #_ %2 <eqw //
95 |* [#tl #Hind * #H * #_ #H2 @Hind % [@H | //]
96 |#a #w1 #tl #Hind * whd in ⊢ ((??%?)→?); #H1 * #H2 #H3 %1
97 @(ex_intro … (a::w1)) @(ex_intro … (flatten S tl)) %
98 [% [@H1 | normalize % /2/] |whd @(ex_intro … tl) /2/]
101 |* [* #w1 * #w2 * * #eqw * #H1 #_ <eqw @cat_to_star //
102 | whd in ⊢ (%→?); #H <H //
107 lemma star_epsilon: ∀S:DeqSet.∀A:word S → Prop.
112 lemma epsilon_cat_r: ∀S.∀A:word S →Prop.
115 [* #w1 * #w2 * * #eqw #inw1 normalize #eqw2 <eqw //
116 |#inA @(ex_intro … w) @(ex_intro … [ ]) /3/
120 lemma epsilon_cat_l: ∀S.∀A:word S →Prop.
123 [* #w1 * #w2 * * #eqw normalize #eqw2 <eqw <eqw2 //
124 |#inA @(ex_intro … ϵ) @(ex_intro … w) /3/
128 lemma distr_cat_r_eps: ∀S.∀A,C:word S →Prop.
129 (A ∪ {ϵ}) · C =1 A · C ∪ C.
130 #S #A #C @eqP_trans [|@distr_cat_r |@eqP_union_l @epsilon_cat_l]