3 (**************************************************************************)
6 (* ||A|| A project by Andrea Asperti *)
8 (* ||I|| Developers: *)
9 (* ||T|| The HELM team. *)
10 (* ||A|| http://helm.cs.unibo.it *)
12 (* \ / This file is distributed under the terms of the *)
13 (* v GNU General Public License Version 2 *)
15 (**************************************************************************)
17 include "arithmetics/nat.ma".
18 include "basics/lists/list.ma".
19 include "basics/sets.ma".
20 include "basics/deqsets.ma".
22 definition word ≝ λS:DeqSet.list S.
24 notation "ϵ" non associative with precedence 90 for @{ 'epsilon }.
25 interpretation "epsilon" 'epsilon = (nil ?).
28 definition cat : ∀S,l1,l2,w.Prop ≝
29 λS.λl1,l2.λw:word S.∃w1,w2.w1 @ w2 = w ∧ l1 w1 ∧ l2 w2.
30 notation "a · b" non associative with precedence 65 for @{ 'middot $a $b}.
31 interpretation "cat lang" 'middot a b = (cat ? a b).
33 let rec flatten (S : DeqSet) (l : list (word S)) on l : word S ≝
34 match l with [ nil ⇒ [ ] | cons w tl ⇒ w @ flatten ? tl ].
36 let rec conjunct (S : DeqSet) (l : list (word S)) (r : word S → Prop) on l: Prop ≝
37 match l with [ nil ⇒ True | cons w tl ⇒ r w ∧ conjunct ? tl r ].
40 definition star ≝ λS.λl.λw:word S.∃lw.flatten ? lw = w ∧ conjunct ? lw l.
41 notation "a ^ *" non associative with precedence 90 for @{ 'star $a}.
42 interpretation "star lang" 'star l = (star ? l).
44 lemma cat_ext_l: ∀S.∀A,B,C:word S →Prop.
45 A ≐ C → A · B ≐ C · B.
46 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
50 lemma cat_ext_r: ∀S.∀A,B,C:word S →Prop.
51 B ≐ C → A · B ≐ A · C.
52 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
56 lemma cat_empty_l: ∀S.∀A:word S→Prop. ∅ · A ≐ ∅.
57 #S #A #w % [|*] * #w1 * #w2 * * #_ *
60 lemma distr_cat_r: ∀S.∀A,B,C:word S →Prop.
61 (A ∪ B) · C ≐ A · C ∪ B · C.
63 [* #w1 * #w2 * * #eqw * /6/ |* * #w1 * #w2 * * /6/]
68 definition deriv ≝ λS.λA:word S → Prop.λa,w. A (a::w).
70 lemma deriv_middot: ∀S,A,B,a. ¬ A ϵ →
71 deriv S (A·B) a ≐ (deriv S A a) · B.
72 #S #A #B #a #noteps #w normalize %
74 [* #w2 * * #_ #Aeps @False_ind /2/
75 |#b #w2 * #w3 * * whd in ⊢ ((??%?)→?); #H destruct
76 #H #H1 @(ex_intro … w2) @(ex_intro … w3) % // % //
78 |* #w1 * #w2 * * #H #H1 #H2 @(ex_intro … (a::w1))
79 @(ex_intro … w2) % // % normalize //
84 lemma espilon_in_star: ∀S.∀A:word S → Prop.
86 #S #A @(ex_intro … [ ]) normalize /2/
89 lemma cat_to_star:∀S.∀A:word S → Prop.
90 ∀w1,w2. A w1 → A^* w2 → A^* (w1@w2).
91 #S #A #w1 #w2 #Aw * #l * #H #H1 @(ex_intro … (w1::l))
95 lemma fix_star: ∀S.∀A:word S → Prop.
98 [* #l generalize in match w; -w cases l [normalize #w * /2/]
99 #w1 #tl #w * whd in ⊢ ((??%?)→?); #eqw whd in ⊢ (%→?); *
100 #w1A #cw1 %1 @(ex_intro … w1) @(ex_intro … (flatten S tl))
101 % /2/ whd @(ex_intro … tl) /2/
102 |* [2: whd in ⊢ (%→?); #eqw <eqw //]
103 * #w1 * #w2 * * #eqw <eqw @cat_to_star
107 lemma star_fix_eps : ∀S.∀A:word S → Prop.
108 A^* ≐ (A - {ϵ}) · A^* ∪ {ϵ}.
111 [* whd in ⊢ ((??%?)→?); #eqw #_ %2 <eqw //
112 |* [#tl #Hind * #H * #_ #H2 @Hind % [@H | //]
113 |#a #w1 #tl #Hind * whd in ⊢ ((??%?)→?); #H1 * #H2 #H3 %1
114 @(ex_intro … (a::w1)) @(ex_intro … (flatten S tl)) %
115 [% [@H1 | normalize % /2/] |whd @(ex_intro … tl) /2/]
118 |* [* #w1 * #w2 * * #eqw * #H1 #_ <eqw @cat_to_star //
119 | whd in ⊢ (%→?); #H <H //
124 lemma star_epsilon: ∀S:DeqSet.∀A:word S → Prop.
129 lemma epsilon_cat_r: ∀S.∀A:word S →Prop.
132 [* #w1 * #w2 * * #eqw #inw1 normalize #eqw2 <eqw //
133 |#inA @(ex_intro … w) @(ex_intro … [ ]) /3/
137 lemma epsilon_cat_l: ∀S.∀A:word S →Prop.
140 [* #w1 * #w2 * * #eqw normalize #eqw2 <eqw <eqw2 //
141 |#inA @(ex_intro … ϵ) @(ex_intro … w) /3/
145 lemma distr_cat_r_eps: ∀S.∀A,C:word S →Prop.
146 (A ∪ {ϵ}) · C ≐ A · C ∪ C.
147 #S #A #C @eqP_trans [|@distr_cat_r |@eqP_union_l @epsilon_cat_l]