]> matita.cs.unibo.it Git - helm.git/blob - weblib/tutorial/chapter6.ma
cadc0065c3bc06a1b610a5312827fdf9a680ea8b
[helm.git] / weblib / tutorial / chapter6.ma
1 include "tutorial/chapter5.ma".
2
3 definition word ≝ λS:DeqSet.list S.
4
5 notation "ϵ" non associative with precedence 90 for @{ 'epsilon }.
6 interpretation "epsilon" 'epsilon = (nil ?).
7
8 (* concatenation *)
9 definition cat : ∀S,l1,l2,w.Prop ≝ 
10   λS.λl1,l2.λw:word S.∃w1,w2.w1 @ w2 = w ∧ l1 w1 ∧ l2 w2.
11 notation "a · b" non associative with precedence 60 for @{ 'middot $a $b}.
12 interpretation "cat lang" 'middot a b = (cat ? a b).
13
14 let rec flatten (S : DeqSet) (l : list (word S)) on l : word S ≝ 
15 match l with [ nil ⇒ [ ] | cons w tl ⇒ w @ flatten ? tl ].
16
17 let rec conjunct (S : DeqSet) (l : list (word S)) (r : word S → Prop) on l: Prop ≝
18 match l with [ nil ⇒ True | cons w tl ⇒ r w ∧ conjunct ? tl r ]. 
19
20 (* kleene's star *)
21 definition star ≝ λS.λl.λw:word S.∃lw.flatten ? lw = w ∧ conjunct ? lw l. 
22 notation "a ^ *" non associative with precedence 90 for @{ 'star $a}.
23 interpretation "star lang" 'star l = (star ? l).
24
25 lemma cat_ext_l: ∀S.∀A,B,C:word S →Prop. 
26   A =1 C  → A · B =1 C · B.
27 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
28 cases (H w1) /6/ 
29 qed.
30
31 lemma cat_ext_r: ∀S.∀A,B,C:word S →Prop. 
32   B =1 C → A · B =1 A · C.
33 #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2
34 cases (H w2) /6/ 
35 qed.
36   
37 lemma cat_empty_l: ∀S.∀A:word S→Prop. ∅ · A =1 ∅.
38 #S #A #w % [|*] * #w1 * #w2 * * #_ *
39 qed.
40
41 lemma distr_cat_r: ∀S.∀A,B,C:word S →Prop.
42   (A ∪ B) · C =1  A · C ∪ B · C. 
43 #S #A #B #C #w %
44   [* #w1 * #w2 * * #eqw * /6/ |* * #w1 * #w2 * * /6/] 
45 qed.
46
47 (* derivative *)
48
49 definition deriv ≝ λS.λA:word S → Prop.λa,w. A (a::w).
50
51 lemma deriv_middot: ∀S,A,B,a. ¬ A ϵ → 
52   deriv S (A·B) a =1 (deriv S A a) · B.
53 #S #A #B #a #noteps #w normalize %
54   [* #w1 cases w1 
55     [* #w2 * * #_ #Aeps @False_ind /2/
56     |#b #w2 * #w3 * * whd in ⊢ ((??%?)→?); #H destruct
57      #H #H1 @(ex_intro … w2) @(ex_intro … w3) % // % //
58     ]
59   |* #w1 * #w2 * * #H #H1 #H2 @(ex_intro … (a::w1))
60    @(ex_intro … w2) % // % normalize //
61   ]
62 qed. 
63
64 (* star properties *)
65 lemma espilon_in_star: ∀S.∀A:word S → Prop.
66   A^* ϵ.
67 #S #A @(ex_intro … [ ]) normalize /2/
68 qed.
69
70 lemma cat_to_star:∀S.∀A:word S → Prop.
71   ∀w1,w2. A w1 → A^* w2 → A^* (w1@w2).
72 #S #A #w1 #w2 #Aw * #l * #H #H1 @(ex_intro … (w1::l)) 
73 % normalize /2/
74 qed.
75
76 lemma fix_star: ∀S.∀A:word S → Prop. 
77   A^* =1 A · A^* ∪ {ϵ}.
78 #S #A #w %
79   [* #l generalize in match w; -w cases l [normalize #w * /2/]
80    #w1 #tl #w * whd in ⊢ ((??%?)→?); #eqw whd in ⊢ (%→?); *
81    #w1A #cw1 %1 @(ex_intro … w1) @(ex_intro … (flatten S tl))
82    % /2/ whd @(ex_intro … tl) /2/
83   |* [2: whd in ⊢ (%→?); #eqw <eqw //]
84    * #w1 * #w2 * * #eqw <eqw @cat_to_star 
85   ]
86 qed.
87
88 lemma star_fix_eps : ∀S.∀A:word S → Prop.
89   A^* =1 (A - {ϵ}) · A^* ∪ {ϵ}.  
90 #S #A #w %
91   [* #l elim l 
92     [* whd in ⊢ ((??%?)→?); #eqw #_ %2 <eqw // 
93     |* [#tl #Hind * #H * #_ #H2 @Hind % [@H | //]
94        |#a #w1 #tl #Hind * whd in ⊢ ((??%?)→?); #H1 * #H2 #H3 %1 
95         @(ex_intro … (a::w1)) @(ex_intro … (flatten S tl)) %
96          [% [@H1 | normalize % /2/] |whd @(ex_intro … tl) /2/]
97        ]
98     ]
99   |* [* #w1 * #w2 * * #eqw * #H1 #_ <eqw @cat_to_star //
100      | whd in ⊢ (%→?); #H <H //
101      ]
102   ]
103 qed. 
104      
105 lemma star_epsilon: ∀S:DeqSet.∀A:word S → Prop.
106   A^* ∪ {ϵ} =1 A^*.
107 #S #A #w % /2/ * // 
108 qed.
109   
110 lemma epsilon_cat_r: ∀S.∀A:word S →Prop.
111    A · {ϵ} =1  A. 
112 #S #A #w %
113   [* #w1 * #w2 * * #eqw #inw1 normalize #eqw2 <eqw //
114   |#inA @(ex_intro … w) @(ex_intro … [ ]) /3/
115   ]
116 qed.
117
118 lemma epsilon_cat_l: ∀S.∀A:word S →Prop.
119    {ϵ} · A =1  A. 
120 #S #A #w %
121   [* #w1 * #w2 * * #eqw normalize #eqw2 <eqw <eqw2 //
122   |#inA @(ex_intro … ϵ) @(ex_intro … w) /3/
123   ]
124 qed.
125
126 lemma distr_cat_r_eps: ∀S.∀A,C:word S →Prop.
127   (A ∪ {ϵ}) · C =1  A · C ∪ C. 
128 #S #A #C @eqP_trans [|@distr_cat_r |@eqP_union_l @epsilon_cat_l]
129 qed.