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