]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/re/moves.ma
In case paramodulation fails we apply unit equalities.
[helm.git] / matita / matita / lib / re / moves.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "re/re.ma".
16 include "basics/lists/listb.ma".
17
18 let rec move (S: DeqSet) (x:S) (E: pitem S) on E : pre S ≝
19  match E with
20   [ pz ⇒ 〈 `∅, false 〉
21   | pe ⇒ 〈 ϵ, false 〉
22   | ps y ⇒ 〈 `y, false 〉
23   | pp y ⇒ 〈 `y, x == y 〉
24   | po e1 e2 ⇒ (move ? x e1) ⊕ (move ? x e2) 
25   | pc e1 e2 ⇒ (move ? x e1) ⊙ (move ? x e2)
26   | pk e ⇒ (move ? x e)^⊛ ].
27   
28 lemma move_plus: ∀S:DeqSet.∀x:S.∀i1,i2:pitem S.
29   move S x (i1 + i2) = (move ? x i1) ⊕ (move ? x i2).
30 // qed.
31
32 lemma move_cat: ∀S:DeqSet.∀x:S.∀i1,i2:pitem S.
33   move S x (i1 · i2) = (move ? x i1) ⊙ (move ? x i2).
34 // qed.
35
36 lemma move_star: ∀S:DeqSet.∀x:S.∀i:pitem S.
37   move S x i^* = (move ? x i)^⊛.
38 // qed.
39
40 definition pmove ≝ λS:DeqSet.λx:S.λe:pre S. move ? x (\fst e).
41
42 lemma pmove_def : ∀S:DeqSet.∀x:S.∀i:pitem S.∀b. 
43   pmove ? x 〈i,b〉 = move ? x i.
44 // qed.
45
46 lemma eq_to_eq_hd: ∀A.∀l1,l2:list A.∀a,b. 
47   a::l1 = b::l2 → a = b.
48 #A #l1 #l2 #a #b #H destruct //
49 qed. 
50
51 lemma same_kernel: ∀S:DeqSet.∀a:S.∀i:pitem S.
52   |\fst (move ? a i)| = |i|.
53 #S #a #i elim i //
54   [#i1 #i2 >move_cat #H1 #H2 whd in ⊢ (???%); <H1 <H2 //
55   |#i1 #i2 >move_plus #H1 #H2 whd in ⊢ (???%); <H1 <H2 //
56   ]
57 qed.
58
59 theorem move_ok:
60  ∀S:DeqSet.∀a:S.∀i:pitem S.∀w: word S. 
61    \sem{move ? a i} w ↔ \sem{i} (a::w).
62 #S #a #i elim i 
63   [normalize /2/
64   |normalize /2/
65   |normalize /2/
66   |normalize #x #w cases (true_or_false (a==x)) #H >H normalize
67     [>(\P H) % [* // #bot @False_ind //| #H1 destruct /2/]
68     |% [@False_ind |#H1 cases (\Pf H) #H2 @H2 destruct //]
69     ]
70   |#i1 #i2 #HI1 #HI2 #w >(sem_cat S i1 i2) >move_cat
71    @iff_trans[|@sem_odot] >same_kernel >sem_cat_w
72    @iff_trans[||@(iff_or_l … (HI2 w))] @iff_or_r 
73    @iff_trans[||@iff_sym @deriv_middot //]
74    @cat_ext_l @HI1
75   |#i1 #i2 #HI1 #HI2 #w >(sem_plus S i1 i2) >move_plus >sem_plus_w 
76    @iff_trans[|@sem_oplus] 
77    @iff_trans[|@iff_or_l [|@HI2]| @iff_or_r //]
78   |#i1 #HI1 #w >move_star 
79    @iff_trans[|@sem_ostar] >same_kernel >sem_star_w 
80    @iff_trans[||@iff_sym @deriv_middot //]
81    @cat_ext_l @HI1
82   ]
83 qed.
84     
85 notation > "x ↦* E" non associative with precedence 60 for @{moves ? $x $E}.
86 let rec moves (S : DeqSet) w e on w : pre S ≝
87  match w with
88   [ nil ⇒ e
89   | cons x w' ⇒ w' ↦* (move S x (\fst e))].
90
91 lemma moves_empty: ∀S:DeqSet.∀e:pre S. 
92   moves ? [ ] e = e.
93 // qed.
94
95 lemma moves_cons: ∀S:DeqSet.∀a:S.∀w.∀e:pre S. 
96   moves ? (a::w)  e = moves ? w (move S a (\fst e)).
97 // qed.
98
99 lemma not_epsilon_sem: ∀S:DeqSet.∀a:S.∀w: word S. ∀e:pre S. 
100   iff ((a::w) ∈ e) ((a::w) ∈ \fst e).
101 #S #a #w * #i #b cases b normalize 
102   [% /2/ * // #H destruct |% normalize /2/]
103 qed.
104
105 lemma same_kernel_moves: ∀S:DeqSet.∀w.∀e:pre S.
106   |\fst (moves ? w e)| = |\fst e|.
107 #S #w elim w //
108 qed.
109
110 theorem decidable_sem: ∀S:DeqSet.∀w: word S. ∀e:pre S. 
111    (\snd (moves ? w e) = true)  ↔ \sem{e} w.
112 #S #w elim w 
113  [* #i #b >moves_empty cases b % /2/
114  |#a #w1 #Hind #e >moves_cons
115   @iff_trans [||@iff_sym @not_epsilon_sem]
116   @iff_trans [||@move_ok] @Hind
117  ]
118 qed.
119
120 lemma not_true_to_false: ∀b.b≠true → b =false.
121 #b * cases b // #H @False_ind /2/ 
122 qed. 
123
124 theorem equiv_sem: ∀S:DeqSet.∀e1,e2:pre S. 
125   iff (\sem{e1} =1 \sem{e2}) (∀w.\snd (moves ? w e1) = \snd (moves ? w e2)).
126 #S #e1 #e2 % 
127 [#same_sem #w 
128   cut (∀b1,b2. iff (b1 = true) (b2 = true) → (b1 = b2)) 
129     [* * // * #H1 #H2 [@sym_eq @H1 //| @H2 //]]
130   #Hcut @Hcut @iff_trans [|@decidable_sem] 
131   @iff_trans [|@same_sem] @iff_sym @decidable_sem
132 |#H #w1 @iff_trans [||@decidable_sem] <H @iff_sym @decidable_sem]
133 qed.
134
135 lemma moves_left : ∀S,a,w,e. 
136   moves S (w@[a]) e = move S a (\fst (moves S w e)). 
137 #S #a #w elim w // #x #tl #Hind #e >moves_cons >moves_cons //
138 qed.
139
140 definition in_moves ≝ λS:DeqSet.λw.λe:pre S. \snd(w ↦* e).
141
142 (*
143 coinductive equiv (S:DeqSet) : pre S → pre S → Prop ≝
144  mk_equiv:
145   ∀e1,e2: pre S.
146    \snd e1  = \snd e2 →
147     (∀x. equiv S (move ? x (\fst e1)) (move ? x (\fst e2))) →
148      equiv S e1 e2.
149 *)
150
151 let rec beqitem S (i1,i2: pitem S) on i1 ≝ 
152   match i1 with
153   [ pz ⇒ match i2 with [ pz ⇒ true | _ ⇒ false]
154   | pe ⇒ match i2 with [ pe ⇒ true | _ ⇒ false]
155   | ps y1 ⇒ match i2 with [ ps y2 ⇒ y1==y2 | _ ⇒ false]
156   | pp y1 ⇒ match i2 with [ pp y2 ⇒ y1==y2 | _ ⇒ false]
157   | po i11 i12 ⇒ match i2 with 
158     [ po i21 i22 ⇒ beqitem S i11 i21 ∧ beqitem S i12 i22
159     | _ ⇒ false]
160   | pc i11 i12 ⇒ match i2 with 
161     [ pc i21 i22 ⇒ beqitem S i11 i21 ∧ beqitem S i12 i22
162     | _ ⇒ false]
163   | pk i11 ⇒ match i2 with [ pk i21 ⇒ beqitem S i11 i21 | _ ⇒ false]
164   ].
165
166 lemma beqitem_true: ∀S,i1,i2. iff (beqitem S i1 i2 = true) (i1 = i2). 
167 #S #i1 elim i1
168   [#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct
169   |#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct
170   |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct
171     [>(\P H) // | @(\b (refl …))]
172   |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct
173     [>(\P H) // | @(\b (refl …))]
174   |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] %
175    normalize #H destruct 
176     [cases (true_or_false (beqitem S i11 i21)) #H1
177       [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H
178       |>H1 in H; normalize #abs @False_ind /2/
179       ]
180     |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) //
181     ]
182   |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] %
183    normalize #H destruct 
184     [cases (true_or_false (beqitem S i11 i21)) #H1
185       [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H
186       |>H1 in H; normalize #abs @False_ind /2/
187       ]
188     |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) //
189     ]
190   |#i3 #Hind #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i4] %
191    normalize #H destruct 
192     [>(proj1 … (Hind i4) H) // |>(proj2 … (Hind i4) (refl …)) //]
193   ]
194 qed. 
195
196 definition DeqItem ≝ λS.
197   mk_DeqSet (pitem S) (beqitem S) (beqitem_true S).
198   
199 unification hint  0 ≔ S; 
200     X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S)
201 (* ---------------------------------------- *) ⊢ 
202     pitem S ≡ carr X.
203     
204 unification hint  0 ≔ S,i1,i2; 
205     X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S)
206 (* ---------------------------------------- *) ⊢ 
207     beqitem S i1 i2 ≡ eqb X i1 i2.
208
209 definition sons ≝ λS:DeqSet.λl:list S.λp:(pre S)×(pre S). 
210  map ?? (λa.〈move S a (\fst (\fst p)),move S a (\fst (\snd p))〉) l.
211
212 lemma memb_sons: ∀S,l.∀p,q:(pre S)×(pre S). memb ? p (sons ? l q) = true →
213   ∃a.(move ? a (\fst (\fst q)) = \fst p ∧
214       move ? a (\fst (\snd q)) = \snd p).
215 #S #l elim l [#p #q normalize in ⊢ (%→?); #abs @False_ind /2/] 
216 #a #tl #Hind #p #q #H cases (orb_true_l … H) -H
217   [#H @(ex_intro … a) <(proj1 … (eqb_true …)H) /2/
218   |#H @Hind @H
219   ]
220 qed.
221
222 let rec bisim S l n (frontier,visited: list ?) on n ≝
223   match n with 
224   [ O ⇒ 〈false,visited〉 (* assert false *)
225   | S m ⇒ 
226     match frontier with
227     [ nil ⇒ 〈true,visited〉
228     | cons hd tl ⇒
229       if beqb (\snd (\fst hd)) (\snd (\snd hd)) then
230         bisim S l m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) 
231         (sons S l hd)) tl) (hd::visited)
232       else 〈false,visited〉
233     ]
234   ].
235   
236 lemma unfold_bisim: ∀S,l,n.∀frontier,visited: list ?.
237   bisim S l n frontier visited =
238   match n with 
239   [ O ⇒ 〈false,visited〉 (* assert false *)
240   | S m ⇒ 
241     match frontier with
242     [ nil ⇒ 〈true,visited〉
243     | cons hd tl ⇒
244       if beqb (\snd (\fst hd)) (\snd (\snd hd)) then
245         bisim S l m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) 
246           (sons S l hd)) tl) (hd::visited)
247       else 〈false,visited〉
248     ]
249   ].
250 #S #l #n cases n // qed.
251   
252 lemma bisim_never: ∀S,l.∀frontier,visited: list ?.
253   bisim S l O frontier visited = 〈false,visited〉.
254 #frontier #visited >unfold_bisim // 
255 qed.
256
257 lemma bisim_end: ∀Sig,l,m.∀visited: list ?.
258   bisim Sig l (S m) [] visited = 〈true,visited〉.
259 #n #visisted >unfold_bisim // 
260 qed.
261
262 lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: list ?.
263 beqb (\snd (\fst p)) (\snd (\snd p)) = true →
264   bisim Sig l (S m) (p::frontier) visited = 
265   bisim Sig l m (unique_append ? (filter ? (λx.notb(memb ? x (p::visited))) 
266     (sons Sig l p)) frontier) (p::visited).
267 #Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // 
268 qed.
269
270 lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: list ?.
271 beqb (\snd (\fst p)) (\snd (\snd p)) = false →
272   bisim Sig l (S m) (p::frontier) visited = 〈false,visited〉.
273 #Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // 
274 qed.
275  
276 definition visited_inv ≝ λS.λe1,e2:pre S.λvisited: list ?.
277 uniqueb ? visited = true ∧  
278   ∀p. memb ? p visited = true → 
279    (∃w.(moves S w e1 = \fst p) ∧ (moves S w e2 = \snd p)) ∧ 
280    (beqb (\snd (\fst p)) (\snd (\snd p)) = true).
281   
282 definition frontier_inv ≝ λS.λfrontier,visited.
283 uniqueb ? frontier = true ∧ 
284 ∀p:(pre S)×(pre S). memb ? p frontier = true →  
285   memb ? p visited = false ∧
286   ∃p1.((memb ? p1 visited = true) ∧
287    (∃a. move ? a (\fst (\fst p1)) = \fst p ∧ 
288         move ? a (\fst (\snd p1)) = \snd p)).
289
290 (* lemma andb_true: ∀b1,b2:bool. 
291   (b1 ∧ b2) = true → (b1 = true) ∧ (b2 = true).
292 #b1 #b2 cases b1 normalize #H [>H /2/ |@False_ind /2/].
293 qed.
294
295 lemma andb_true_r: ∀b1,b2:bool. 
296   (b1 = true) ∧ (b2 = true) → (b1 ∧ b2) = true.
297 #b1 #b2 cases b1 normalize * // 
298 qed. *)
299
300 lemma notb_eq_true_l: ∀b. notb b = true → b = false.
301 #b cases b normalize //
302 qed.
303
304 (*
305 lemma notb_eq_true_r: ∀b. b = false → notb b = true.
306 #b cases b normalize //
307 qed.
308  
309 lemma notb_eq_false_l:∀b. notb b = false → b = true.
310 #b cases b normalize //
311 qed.
312
313 lemma notb_eq_false_r:∀b. b = true → notb b = false.
314 #b cases b normalize //
315 qed. *)
316
317 (* include "arithmetics/exp.ma". *)
318
319 let rec pos S (i:re S) on i ≝ 
320   match i with
321   [ z ⇒ 0
322   | e ⇒ 0
323   | s y ⇒ 1
324   | o i1 i2 ⇒ pos S i1 + pos S i2
325   | c i1 i2 ⇒ pos S i1 + pos S i2
326   | k i ⇒ pos S i
327   ].
328
329   
330 let rec pitem_enum S (i:re S) on i ≝
331   match i with
332   [ z ⇒ [pz S]
333   | e ⇒ [pe S]
334   | s y ⇒ [ps S y; pp S y]
335   | o i1 i2 ⇒ compose ??? (po S) (pitem_enum S i1) (pitem_enum S i2)
336   | c i1 i2 ⇒ compose ??? (pc S) (pitem_enum S i1) (pitem_enum S i2)
337   | k i ⇒ map ?? (pk S) (pitem_enum S i)
338   ].
339   
340 lemma pitem_enum_complete : ∀S.∀i:pitem S.
341   memb (DeqItem S) i (pitem_enum S (|i|)) = true.
342 #S #i elim i 
343   [1,2://
344   |3,4:#c normalize >(\b (refl … c)) //
345   |5,6:#i1 #i2 #Hind1 #Hind2 @(memb_compose (DeqItem S) (DeqItem S)) //
346   |#i #Hind @(memb_map (DeqItem S)) //
347   ]
348 qed.
349
350 definition pre_enum ≝ λS.λi:re S.
351   compose ??? (λi,b.〈i,b〉) (pitem_enum S i) [true;false].
352   
353 lemma pre_enum_complete : ∀S.∀e:pre S.
354   memb ? e (pre_enum S (|\fst e|)) = true.
355 #S * #i #b @(memb_compose (DeqItem S) DeqBool ? (λi,b.〈i,b〉))
356 // cases b normalize //
357 qed.
358  
359 definition space_enum ≝ λS.λi1,i2:re S.
360   compose ??? (λe1,e2.〈e1,e2〉) (pre_enum S i1) (pre_enum S i2).
361
362 lemma space_enum_complete : ∀S.∀e1,e2: pre S.
363   memb ? 〈e1,e2〉 (space_enum S (|\fst e1|) (|\fst e2|)) = true.
364 #S #e1 #e2 @(memb_compose … (λi,b.〈i,b〉))
365 // qed.
366
367 definition visited_inv_1 ≝ λS.λe1,e2:pre S.λvisited: list ?.
368 uniqueb ? visited = true ∧  
369   ∀p. memb ? p visited = true → 
370     ∃w.(moves S w e1 = \fst p) ∧ (moves S w e2 = \snd p). 
371    
372 lemma bisim_ok1: ∀S.∀e1,e2:pre S.\sem{e1}=1\sem{e2} → 
373  ∀l,n.∀frontier,visited:list (*(space S) *) ((pre S)×(pre S)).
374  |space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|→
375  visited_inv_1 S e1 e2 visited →  frontier_inv S frontier visited →
376  \fst (bisim S l n frontier visited) = true.
377 #Sig #e1 #e2 #same #l #n elim n 
378   [#frontier #visited #abs * #unique #H @False_ind @(absurd … abs)
379    @le_to_not_lt @sublist_length // * #e11 #e21 #membp 
380    cut ((|\fst e11| = |\fst e1|) ∧ (|\fst e21| = |\fst e2|))
381    [|* #H1 #H2 <H1 <H2 @space_enum_complete]
382    cases (H … membp) #w * #we1 #we2 <we1 <we2 % >same_kernel_moves //    
383   |#m #HI * [#visited #vinv #finv >bisim_end //]
384    #p #front_tl #visited #Hn * #u_visited #vinv * #u_frontier #finv
385    cases (finv p (memb_hd …)) #Hp * #p2 * #visited_p2
386    * #a * #movea1 #movea2
387    cut (∃w.(moves Sig w e1 = \fst p) ∧ (moves Sig w e2 = \snd p))
388      [cases (vinv … visited_p2) -vinv #w1 * #mw1 #mw2 
389       @(ex_intro … (w1@[a])) % //] 
390    -movea2 -movea1 -a -visited_p2 -p2 #reachp
391    cut (beqb (\snd (\fst p)) (\snd (\snd p)) = true)
392      [cases reachp #w * #move_e1 #move_e2 <move_e1 <move_e2
393       @(\b ?) @(proj1 … (equiv_sem … )) @same] #ptest 
394    >(bisim_step_true … ptest) @HI -HI 
395      [<plus_n_Sm //
396      |% [whd in ⊢ (??%?); >Hp whd in ⊢ (??%?); //]
397        #p1 #H (cases (orb_true_l … H))
398          [#eqp <(\P eqp) // 
399          |#visited_p1 @(vinv … visited_p1)
400          ]
401      |whd % [@unique_append_unique @(andb_true_r … u_frontier)]
402       @unique_append_elim #q #H
403        [% 
404          [@notb_eq_true_l @(filter_true … H) 
405          |@(ex_intro … p) % [@memb_hd|@(memb_sons … (memb_filter_memb … H))]
406          ]
407        |cases (finv q ?) [|@memb_cons //]
408         #nvq * #p1 * #Hp1 #reach %
409          [cut ((p==q) = false) [|#Hpq whd in ⊢ (??%?); >Hpq @nvq]
410           cases (andb_true … u_frontier) #notp #_ 
411           @(not_memb_to_not_eq … H) @notb_eq_true_l @notp 
412          |cases (proj2 … (finv q ?)) 
413            [#p1 *  #Hp1 #reach @(ex_intro … p1) % // @memb_cons //
414            |@memb_cons //
415            ]
416         ]
417       ]  
418     ]
419   ]
420 qed.
421
422 definition all_true ≝ λS.λl.∀p:(pre S) × (pre S). memb ? p l = true → 
423   (beqb (\snd (\fst p)) (\snd (\snd p)) = true).
424
425 definition sub_sons ≝ λS,l,l1,l2.∀x:(pre S) × (pre S).∀a:S. 
426 memb ? x l1 = true → memb S a l = true →
427   memb ? 〈move ? a (\fst (\fst x)), move ? a (\fst (\snd x))〉 l2 = true.
428
429 lemma reachable_bisim: 
430  ∀S,l,n.∀frontier,visited,visited_res:list ?.
431  all_true S visited →
432  sub_sons S l visited (frontier@visited) →
433  bisim S l n frontier visited = 〈true,visited_res〉 → 
434   (sub_sons S l visited_res visited_res ∧ 
435    sublist ? visited visited_res ∧
436    all_true S visited_res).
437 #S #l #n elim n
438   [#fron #vis #vis_res #_ #_ >bisim_never #H destruct
439   |#m #Hind * 
440     [(* case empty frontier *)
441      -Hind #vis #vis_res #allv #H normalize in  ⊢ (%→?);
442      #H1 destruct % // % // #p /2 by / 
443     |#hd cases (true_or_false (beqb (\snd (\fst hd)) (\snd (\snd hd))))
444       [|(* case head of the frontier is non ok (absurd) *)
445        #H #tl #vis #vis_res #allv >(bisim_step_false … H) #_ #H1 destruct]
446      (* frontier = hd:: tl and hd is ok *)
447      #H #tl #visited #visited_res #allv >(bisim_step_true … H)
448      (* new_visited = hd::visited are all ok *)
449      cut (all_true S (hd::visited)) 
450       [#p #H1 cases (orb_true_l … H1) [#eqp <(\P eqp) @H |@allv]]
451      (* we now exploit the induction hypothesis *)
452      #allh #subH #bisim cases (Hind … allh … bisim) -Hind
453       [* #H1 #H2 #H3 % // % // #p #H4 @H2 @memb_cons //]
454      (* the only thing left to prove is the sub_sons invariant *)  
455      #x #a #membx #memba
456      cases (orb_true_l … membx)
457       [(* case x = hd *) 
458        #eqhdx >(proj1 … (eqb_true …) eqhdx)
459        (* xa is the son of x w.r.t. a; we must distinguish the case xa 
460         was already visited form the case xa is new *)
461        letin xa ≝ 〈move S a (\fst (\fst x)), move S a (\fst (\snd x))〉
462        cases (true_or_false … (memb ? xa (x::visited)))
463         [(* xa visited - trivial *) #membxa @memb_append_l2 //
464         |(* xa new *) #membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l
465           [>membxa //
466           |(* this can be probably improved *)
467            generalize in match memba; -memba elim l
468             [whd in ⊢ (??%?→?); #abs @False_ind /2/
469             |#b #others #Hind #memba cases (orb_true_l … memba) #H
470               [>(proj1 … (eqb_true …) H) @memb_hd
471               |@memb_cons @Hind //
472               ]
473             ]
474           ]
475         ]
476       |(* case x in visited *)
477        #H1 letin xa ≝ 〈move S a (\fst (\fst x)), move S a (\fst (\snd x))〉
478        cases (memb_append … (subH x a H1 memba))  
479         [#H2 (cases (orb_true_l … H2)) 
480           [#H3 @memb_append_l2 >(proj1 … (eqb_true …) H3) @memb_hd
481           |#H3 @memb_append_l1 @sublist_unique_append_l2 @H3
482           ]
483         |#H2 @memb_append_l2 @memb_cons @H2
484         ]
485       ]
486     ]
487   ]
488 qed.
489
490 (* pit state *)
491 let rec blank_item (S: DeqSet) (i: re S) on i :pitem S ≝
492  match i with
493   [ z ⇒ `∅
494   | e ⇒ ϵ
495   | s y ⇒ `y
496   | o e1 e2 ⇒ (blank_item S e1) + (blank_item S e2) 
497   | c e1 e2 ⇒ (blank_item S e1) · (blank_item S e2)
498   | k e ⇒ (blank_item S e)^* ].
499  
500 definition pit_pre ≝ λS.λi.〈blank_item S (|i|), false〉. 
501
502 let rec occur (S: DeqSet) (i: re S) on i ≝  
503   match i with
504   [ z ⇒ [ ]
505   | e ⇒ [ ]
506   | s y ⇒ [y]
507   | o e1 e2 ⇒ unique_append ? (occur S e1) (occur S e2) 
508   | c e1 e2 ⇒ unique_append ? (occur S e1) (occur S e2) 
509   | k e ⇒ occur S e].
510   
511 axiom memb_single: ∀S,a,x. memb S a [x] = true → a = x.
512
513 axiom tech: ∀b. b ≠ true → b = false.
514 axiom tech2: ∀b. b = false → b ≠ true.
515
516 lemma not_occur_to_pit: ∀S,a.∀i:pitem S. memb S a (occur S (|i|)) = false →
517   move S a i  = pit_pre S i.
518 #S #a #i elim i //
519   [#x cases (true_or_false (a==x)) 
520     [#H >(proj1 …(eqb_true …) H) whd in ⊢ ((??%?)→?); 
521      >(proj2 …(eqb_true …) (refl …)) whd in ⊢ ((??%?)→?); #abs @False_ind /2/
522     |#H normalize >H //
523     ]
524   |#i1 #i2 #Hind1 #Hind2 #H >move_cat >Hind1 [2:@tech 
525    @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l1 //]
526    >Hind2 [2:@tech @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l2 //]
527    //
528   |#i1 #i2 #Hind1 #Hind2 #H >move_plus >Hind1 [2:@tech 
529    @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l1 //]
530    >Hind2 [2:@tech @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l2 //]
531    //
532   |#i #Hind #H >move_star >Hind // @H
533   ]
534 qed.
535
536 lemma move_pit: ∀S,a,i. move S a (\fst (pit_pre S i)) = pit_pre S i.
537 #S #a #i elim i //
538   [#i1 #i2 #Hind1 #Hind2 >move_cat >Hind1 >Hind2 // 
539   |#i1 #i2 #Hind1 #Hind2 >move_plus >Hind1 >Hind2 // 
540   |#i #Hind >move_star >Hind //
541   ]
542 qed. 
543
544 lemma moves_pit: ∀S,w,i. moves S w (pit_pre S i) = pit_pre S i.
545 #S #w #i elim w // #a #tl >moves_cons // 
546 qed. 
547  
548 lemma to_pit: ∀S,w,e. ¬ sublist S w (occur S (|\fst e|)) →
549  moves S w e = pit_pre S (\fst e).
550 #S #w elim w
551   [(* orribile *)
552    #e * #H @False_ind @H normalize #a #abs @False_ind /2/
553   |#a #tl #Hind #e #H cases (true_or_false (memb S a (occur S (|\fst e|))))
554     [#Htrue >moves_cons whd in ⊢ (???%); <(same_kernel … a) 
555      @Hind >same_kernel @(not_to_not … H) #H1 #b #memb cases (orb_true_l … memb)
556       [#H2 <(proj1 … (eqb_true …) H2) // |#H2 @H1 //]
557     |#Hfalse >moves_cons >not_occur_to_pit //
558     ]
559   ]
560 qed.
561     
562 definition occ ≝ λS.λe1,e2:pre S. 
563   unique_append ? (occur S (|\fst e1|)) (occur S (|\fst e2|)).
564
565 (* definition occS ≝ λS:DeqSet.λoccur.
566   PSig S (λx.memb S x occur = true). *)
567
568 lemma occ_enough: ∀S.∀e1,e2:pre S.
569 (∀w.(sublist S w (occ S e1 e2))→
570  (beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true) \to
571 ∀w.(beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true.
572 #S #e1 #e2 #H #w
573 cut (sublist S w (occ S e1 e2) ∨ ¬(sublist S w (occ S e1 e2)))
574 [elim w 
575   [%1 #a normalize in ⊢ (%→?); #abs @False_ind /2/
576   |#a #tl * #subtl 
577     [cases (true_or_false (memb S a (occ S e1 e2))) #memba
578       [%1 whd #x #membx cases (orb_true_l … membx)
579         [#eqax <(proj1 … (eqb_true …) eqax) //
580         |@subtl
581         ]
582       |%2 @(not_to_not … (tech2 … memba)) #H1 @H1 @memb_hd
583       ]
584     |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons //
585     ] 
586   ]
587 |* [@H] 
588  #H >to_pit 
589   [2: @(not_to_not … H) #H1 #a #memba  @sublist_unique_append_l1 @H1 //]
590  >to_pit
591   [2: @(not_to_not … H) #H1 #a #memba  @sublist_unique_append_l2 @H1 //]
592  //
593 ]
594 qed.
595
596 lemma bisim_char: ∀S.∀e1,e2:pre S.
597 (∀w.(beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true) → 
598 \sem{e1}=1\sem{e2}.
599 #S #e1 #e2 #H @(proj2 … (equiv_sem …)) #w @(\P ?) @H
600 qed.
601
602 lemma bisim_ok2: ∀S.∀e1,e2:pre S.
603  (beqb (\snd e1) (\snd e2) = true) → ∀n.
604  \fst (bisim S (occ S e1 e2) n (sons S (occ S e1 e2) 〈e1,e2〉) [〈e1,e2〉]) = true → 
605    \sem{e1}=1\sem{e2}.
606 #S #e1 #e2 #Hnil #n 
607 letin rsig ≝ (occ S e1 e2)
608 letin frontier ≝ (sons S rsig 〈e1,e2〉)
609 letin visited_res ≝ (\snd (bisim S rsig n frontier [〈e1,e2〉])) 
610 #bisim_true
611 cut (bisim S rsig n frontier [〈e1,e2〉] = 〈true,visited_res〉)
612   [<bisim_true <eq_pair_fst_snd //] #H
613 cut (all_true S [〈e1,e2〉]) 
614   [#p #Hp cases (orb_true_l … Hp) 
615     [#eqp <(proj1 … (eqb_true …) eqp) // 
616     | whd in ⊢ ((??%?)→?); #abs @False_ind /2/
617     ]] #allH 
618 cut (sub_sons S rsig [〈e1,e2〉] (frontier@[〈e1,e2〉]))
619   [#x #a #H1 cases (orb_true_l … H1) 
620     [#eqx <(proj1 … (eqb_true …) eqx) #H2 @memb_append_l1 
621      whd in ⊢ (??(???%)?); @(memb_map … H2)
622     |whd in ⊢ ((??%?)→?); #abs @False_ind /2/
623     ]
624   ] #init
625 cases (reachable_bisim … allH init … H) * #H1 #H2 #H3
626 cut (∀w.sublist ? w (occ S e1 e2)→∀p.memb ? p visited_res = true → 
627   memb ? 〈moves ? w (\fst p), moves ? w (\snd p)〉 visited_res = true)
628   [#w elim w  [#_ #p #H4 >moves_empty >moves_empty <eq_pair_fst_snd //] 
629    #a #w1 #Hind #Hsub * #e11 #e21 #visp >moves_cons >moves_cons 
630    @(Hind ? 〈?,?〉) [#x #H4 @Hsub @memb_cons //] 
631    @(H1 〈?,?〉) [@visp| @Hsub @memb_hd]] #all_reach
632 @bisim_char @occ_enough
633 #w #Hsub @(H3 〈?,?〉) @(all_reach w Hsub 〈?,?〉) @H2 //
634 qed.
635   
636 (*
637 definition tt ≝ ps Bin true.
638 definition ff ≝ ps Bin false.
639 definition eps ≝ pe Bin.
640 definition exp1 ≝ (ff + tt · ff).
641 definition exp2 ≝  ff · (eps + tt).
642
643 definition exp3 ≝ move Bin true (\fst (•exp1)).
644 definition exp4 ≝ move Bin true (\fst (•exp2)).
645 definition exp5 ≝ move Bin false (\fst (•exp1)).
646 definition exp6 ≝ move Bin false (\fst (•exp2)). *)
647
648
649
650