]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/match.ma
Closed some daemons
[helm.git] / matita / matita / lib / turing / multi_universal / match.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 "turing/multi_universal/compare.ma".
16 include "turing/multi_universal/par_test.ma".
17 include "turing/multi_universal/moves_2.ma".
18
19 definition Rtc_multi_true ≝ 
20   λalpha,test,n,i.λt1,t2:Vector ? (S n).
21    (∃c. current alpha (nth i ? t1 (niltape ?)) = Some ? c ∧ test c = true) ∧ t2 = t1.
22    
23 definition Rtc_multi_false ≝ 
24   λalpha,test,n,i.λt1,t2:Vector ? (S n).
25     (∀c. current alpha (nth i ? t1 (niltape ?)) = Some ? c → test c = false) ∧ t2 = t1.
26
27 lemma sem_test_char_multi :
28   ∀alpha,test,n,i.i ≤ n → 
29   inject_TM ? (test_char ? test) n i ⊨ 
30   [ tc_true : Rtc_multi_true alpha test n i, Rtc_multi_false alpha test n i ].
31 #alpha #test #n #i #Hin #int
32 cases (acc_sem_inject … Hin (sem_test_char alpha test) int)
33 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
34 [ @Hloop
35 | #Hqtrue lapply (Htrue Hqtrue) * * * #c *
36   #Hcur #Htestc #Hnth_i #Hnth_j %
37   [ %{c} % //
38   | @(eq_vec … (niltape ?)) #i0 #Hi0
39     cases (decidable_eq_nat i0 i) #Hi0i
40     [ >Hi0i @Hnth_i
41     | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
42 | #Hqfalse lapply (Hfalse Hqfalse) * * #Htestc #Hnth_i #Hnth_j %
43   [ @Htestc
44   | @(eq_vec … (niltape ?)) #i0 #Hi0
45     cases (decidable_eq_nat i0 i) #Hi0i
46     [ >Hi0i @Hnth_i
47     | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
48 qed.
49
50 definition Rm_test_null_true ≝ 
51   λalpha,n,i.λt1,t2:Vector ? (S n).
52    current alpha (nth i ? t1 (niltape ?)) ≠ None ? ∧ t2 = t1.
53    
54 definition Rm_test_null_false ≝ 
55   λalpha,n,i.λt1,t2:Vector ? (S n).
56     current alpha (nth i ? t1 (niltape ?)) = None ? ∧ t2 = t1.
57
58 lemma sem_test_null_multi : ∀alpha,n,i.i ≤ n → 
59   inject_TM ? (test_null ?) n i ⊨ 
60     [ tc_true : Rm_test_null_true alpha n i, Rm_test_null_false alpha n i ].
61 #alpha #n #i #Hin #int
62 cases (acc_sem_inject … Hin (sem_test_null alpha) int)
63 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
64 [ @Hloop
65 | #Hqtrue lapply (Htrue Hqtrue) * * #Hcur #Hnth_i #Hnth_j % //
66   @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) #Hi0i
67   [ >Hi0i @sym_eq @Hnth_i | @sym_eq @Hnth_j @sym_not_eq // ] ] 
68 | #Hqfalse lapply (Hfalse Hqfalse) * * #Hcur #Hnth_i #Hnth_j %
69   [ @Hcur
70   | @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) // 
71     #Hi0i @sym_eq @Hnth_j @sym_not_eq // ] ] 
72 qed.
73
74 definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
75   match (nth src (option sig) v (None ?)) with 
76   [ None ⇒  false 
77   | Some x ⇒  notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
78
79 definition mmove_states ≝ initN 2.
80
81 definition mmove0 : mmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
82 definition mmove1 : mmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
83
84 definition trans_mmove ≝ 
85  λi,sig,n,D.
86  λp:mmove_states × (Vector (option sig) (S n)).
87  let 〈q,a〉 ≝ p in match (pi1 … q) with
88  [ O ⇒ 〈mmove1,change_vec ? (S n) (null_action ? n) (〈None ?,D〉) i〉
89  | S _ ⇒ 〈mmove1,null_action sig n〉 ].
90
91 definition mmove ≝ 
92   λi,sig,n,D.
93   mk_mTM sig n mmove_states (trans_mmove i sig n D) 
94     mmove0 (λq.q == mmove1).
95     
96 definition Rm_multi ≝ 
97   λalpha,n,i,D.λt1,t2:Vector ? (S n).
98   t2 = change_vec ? (S n) t1 (tape_move alpha (nth i ? t1 (niltape ?)) D) i.
99    
100 lemma sem_move_multi :
101   ∀alpha,n,i,D.i ≤ n → 
102   mmove i alpha n D ⊨ Rm_multi alpha n i D.
103 #alpha #n #i #D #Hin #int %{2}
104 %{(mk_mconfig ? mmove_states n mmove1 ?)} 
105 [| %
106  [ whd in ⊢ (??%?); @eq_f whd in ⊢ (??%?); @eq_f %
107  | whd >tape_move_multi_def
108    <(change_vec_same … (ctapes …) i (niltape ?))
109    >pmap_change <tape_move_multi_def >tape_move_null_action % ] ]
110  qed.
111   
112 definition rewind ≝ λsrc,dst,sig,n.
113   parmove src dst sig n L · mmove src sig n R · mmove dst sig n R.
114
115 definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
116   (∀x,x0,xs,rs.
117     nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
118     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
119     nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
120     outt = change_vec ?? 
121            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
122            (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧
123   (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → 
124    ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → 
125    outt = int).
126
127 (*
128 theorem accRealize_to_Realize :
129   ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
130   M ⊨ [ acc: Rtrue, Rfalse ] →  M ⊨ Rtrue ∪ Rfalse.
131 #sig #n #M #Rtrue #Rfalse #acc #HR #t
132 cases (HR t) #k * #outc * * #Hloop
133 #Htrue #Hfalse %{k} %{outc} % // 
134 cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
135 [ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
136 qed.
137 *)
138
139 lemma sem_rewind : ∀src,dst,sig,n.
140   src ≠ dst → src < S n → dst < S n → 
141   rewind src dst sig n ⊨ R_rewind src dst sig n.
142 #src #dst #sig #n #Hneq #Hsrc #Hdst
143 @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst) ?)
144 [| @(sem_seq_app sig n ????? (sem_move_multi … R ?) (sem_move_multi … R ?)) //
145  @le_S_S_to_le // ]
146 #ta #tb * #tc * * #Htc #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb %
147 [ #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst
148   >(Htc ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd;
149   [|>Hmidta_dst //
150   |>length_append >length_append >Hlen % ]
151   >change_vec_commute [|@sym_not_eq //]
152   >change_vec_change_vec
153   >nth_change_vec_neq [|@sym_not_eq //]
154   >nth_change_vec // >reverse_append >reverse_single
155   >reverse_append >reverse_single normalize in match (tape_move ???);
156   >rev_append_def >append_nil #Htd >Htd in Htb;
157   >change_vec_change_vec >nth_change_vec //
158   cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); //
159 | #x #rs #Hmidta_src #ls0 #y #rs0 #Hmidta_dst
160   lapply (Htc … Hmidta_src … (refl ??) Hmidta_dst) -Htc #Htc >Htc in Htd;
161   >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
162   >nth_change_vec_neq [|@sym_not_eq //] 
163   >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%);
164   [ #Hls0 #Htd >Htd in Htb; 
165     >nth_change_vec // >change_vec_change_vec 
166     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
167     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
168   | #l1 #ls1 #Hls0 #Htd >Htd in Htb;
169     >nth_change_vec // >change_vec_change_vec 
170     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
171     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
172 ]]
173 qed.
174
175 definition match_step ≝ λsrc,dst,sig,n.
176   compare src dst sig n ·
177      (ifTM ?? (partest sig n (match_test src dst sig ?))
178       (single_finalTM ??
179         (rewind src dst sig n · (inject_TM ? (move_r ?) n dst)))
180       (nop …)
181       partest1).
182
183 (* we assume the src is a midtape
184    we stop
185    if the dst is out of bounds (outt = int)
186    or dst.right is shorter than src.right (outt.current → None)
187    or src.right is a prefix of dst.right (out = just right of the common prefix) *)
188 definition R_match_step_false ≝ 
189   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
190   ∀ls,x,xs.
191   nth src ? int (niltape ?) = midtape sig ls x xs →
192   ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
193     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
194       xs = rs0@xs0 ∧
195       current sig (nth dst (tape sig) outt (niltape sig)) = None ?) ∨
196     (∃ls0,rs0. 
197      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
198      (* ∀rsj,c. 
199      rs0 = c::rsj → *)
200      outt = change_vec ??
201             (change_vec ?? int (mk_tape sig (reverse ? xs@x::ls) (None ?) [ ]) src)
202             (mk_tape sig (reverse ? xs@x::ls0) (option_hd ? rs0) (tail ? rs0)) dst).
203
204 (*
205    we assume the src is a midtape [ ] s rs
206    if we iterate
207    then dst.current = Some ? s1
208    and  if s ≠ s1 then outt = int.dst.move_right()
209    and  if s = s1 
210         then int.src.right and int.dst.right have a common prefix
211         and  the heads of their suffixes are different
212         and  outt = int.dst.move_right().
213                
214  *)
215 definition R_match_step_true ≝ 
216   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
217   ∀s,rs.nth src ? int (niltape ?) = midtape ? [ ] s rs → 
218   outt = change_vec ?? int 
219          (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst ∧
220   (current sig (nth dst (tape sig) int (niltape sig)) = Some ? s → 
221    ∃xs,ci,rs',ls0,cj,rs0.
222    rs = xs@ci::rs' ∧
223    nth dst ? int (niltape ?) = midtape sig ls0 s (xs@cj::rs0) ∧
224    ci ≠ cj).
225   
226 axiom daemon : ∀X:Prop.X.
227       
228 lemma sem_match_step :
229   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
230   match_step src dst sig n ⊨ 
231     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
232       R_match_step_true src dst sig n, 
233       R_match_step_false src dst sig n ].
234 #src #dst #sig #n #Hneq #Hsrc #Hdst 
235 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
236     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
237       (sem_seq … 
238         (sem_rewind ???? Hneq Hsrc Hdst) 
239         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
240       (sem_nop …)))
241 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
242   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
243   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
244     whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
245     >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
246   | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
247     cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
248     [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
249       whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
250       >nth_current_chars >nth_current_chars >Hs >Hcurta_dst 
251       normalize in ⊢ (%→?); #H destruct (H)
252     | #s0 #Hs0
253       cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
254       cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
255       cases (true_or_false (s == s0)) #Hss0
256       [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0) 
257         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
258         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
259           >nth_current_chars >nth_current_chars >Hcurtc_dst 
260           cases (current ? (nth src …)) [| #x ] 
261           normalize in ⊢ (%→?); #H destruct (H) 
262         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
263           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
264           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
265             [>nth_change_vec [cases (append ???) // | @Hsrc] 
266             |@(not_to_not … Hneq) //
267             ]]
268           normalize in ⊢ (%→?); #H destruct (H) ]
269         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
270           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * *
271           >Htc >change_vec_commute // >nth_change_vec //
272           >change_vec_commute [|@sym_not_eq //] >nth_change_vec // #Hte #_ #Htb
273           #s' #rs' >Hmidta_src #H destruct (H)
274           lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
275           >change_vec_commute // >change_vec_change_vec
276           >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
277           >Hte in Htb; * * #_ >nth_change_vec // #Htb1
278           lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 %
279           [ @(eq_vec … (niltape ?)) #i #Hi
280             cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
281             [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst
282               >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] %
283             | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)]
284               >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%); 
285               <Hrs <Hmidta_src [|@(\Pf Hdsti)] >change_vec_same % ]
286           | #_ >Hmidta_dst >Hrs0
287             %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
288           ]
289         ]
290       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
291         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
292         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
293         #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2
294         #s1 #rs1 >Hmidta_src #H destruct (H)
295         lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
296         [ @(eq_vec … (niltape ?)) #i #Hi
297           cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
298           [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
299             cases rs0 [|#r2 #rs2] %
300           | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
301         | >Hs0 #H destruct (H) @False_ind cases (Hss0) /2/ ]
302       ]
303     ]
304   ]
305 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
306   whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
307   lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
308   cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
309   [ #Hcurta_dst % % % // @Hcomp1 %2 //
310   | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
311     #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
312     [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
313     | >(?:tc=ta) in Htest; 
314       [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
315         #Hxx0' destruct (Hxx0') % ]
316       whd in ⊢ (??%?→?); 
317       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
318       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
319       cases (Hcomp2 … Hta_src Hta_dst) [ *
320       [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} % // % //
321       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
322       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
323         #Hci #Hxs #Hrs0 #Htc @False_ind
324         whd in Htest:(??%?); 
325         >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest; 
326         [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
327           >nth_change_vec //]
328         >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj) 
329         [|>nth_current_chars >Htc >nth_change_vec //]
330         normalize #H destruct (H) ] ] ]
331 qed.
332
333 definition match_m ≝ λsrc,dst,sig,n.
334   whileTM … (match_step src dst sig n) 
335     (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
336
337 definition R_match_m ≝ 
338   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
339   ∀x,rs.
340   nth src ? int (niltape ?) = midtape sig [ ] x rs →
341   (current sig (nth dst (tape sig) int (niltape sig)) = None ? → outt = int) ∧
342   (∀ls0,x0,rs0.
343    nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
344    (∃l,l1.x0::rs0 = l@x::rs@l1 ∧
345     outt = change_vec ?? 
346            (change_vec ?? int 
347              (mk_tape sig (reverse ? rs@[x]) (None ?) [ ]) src)
348            (mk_tape sig ((reverse ? (l@x::rs))@ls0) (option_hd ? l1) (tail ? l1)) dst) ∨
349     ∀l,l1.x0::rs0 ≠ l@x::rs@l1).
350
351 lemma not_sub_list_merge : 
352   ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
353 #T #a #b #H1 #H2 #l elim l normalize //
354 qed.
355
356 lemma not_sub_list_merge_2 : 
357   ∀T:DeqSet.∀a,b:list T.∀t. (∀l1.t::a ≠ b@l1) → (∀l,l1.a ≠ l@b@l1) → ∀l,l1.t::a ≠ l@b@l1.
358 #T #a #b #t #H1 #H2 #l elim l //
359 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
360 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
361 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
362 qed.
363
364
365 lemma wsem_match_m : ∀src,dst,sig,n.
366 src ≠ dst → src < S n → dst < S n → 
367   match_m src dst sig n ⊫ R_match_m src dst sig n.
368 #src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
369 lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
370 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
371 [ #Hfalse #x #xs #Hmid_src
372   cases (Hfalse … Hmid_src) -Hfalse 
373   [(* current dest = None *) *
374     [ * #Hcur_dst #Houtc %
375       [#_ >Houtc //
376       | #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst; 
377         normalize in ⊢ (%→?); #H destruct (H)
378       ]
379     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
380       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
381       | #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
382         >Hrs0 cases xs0
383         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
384           (* change false case 
385              #cj #ls2 #H destruct (H) *) @daemon
386         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
387           >length_append whd in ⊢ (??%(??%)→?); >length_append
388           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
389           #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
390           >associative_plus >associative_plus 
391           #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
392           #e2 destruct (e2)
393         ]
394       ]
395     ]
396   |* #ls0 * #rs0 * #Hmid_dst #Houtc %
397     [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
398     |#ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
399      %1 %{[ ]} %{rs0} % [%] 
400      >reverse_cons >associative_append >Houtc %
401     ]
402   ]
403 |-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
404  #x #xs #Hmidta_src
405  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
406  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
407   [#Hcurta_dst % 
408     [#_ whd in Htrue; >Hmidta_src in Htrue; #Htrue
409      cases (Htrue ?? (refl ??)) -Htrue >Hcurta_dst
410      (* dovremmo sapere che ta.dst è sul margine destro, da cui la move non 
411         ha effetto *) #_ cut (tc = ta) [@daemon] #Htc destruct (Htc) #_
412      cases (IH … Hmidta_src) #Houtc #_ @Houtc //
413     |#ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst;
414      normalize in ⊢ (%→?); #H destruct (H)
415     ]
416   | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
417     #ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst; normalize in ⊢ (%→?);
418     #H destruct (H) whd in Htrue; >Hmidta_src in Htrue; #Htrue
419     cases (Htrue ?? (refl …)) -Htrue >Hmidta_dst #Htc
420     cases (true_or_false (x==c)) #eqx
421     [ lapply (\P eqx) -eqx #eqx destruct (eqx)
422       #Htrue cases (Htrue (refl ??)) -Htrue
423       #xs0 * #ci * #rs' * #ls1 * #cj * #rs1 * * #Hxs #H destruct (H) #Hcicj
424       >Htc in IH; whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
425       #IH cases (IH … Hmidta_src) -IH #_ >nth_change_vec //
426       cut (∃x1,xs1.xs0@cj::rs1 = x1::xs1)
427       [ cases xs0 [ %{cj} %{rs1} % | #x1 #xs1 %{x1} %{(xs1@cj::rs1)} % ] ] * #x1 * #xs1
428       #Hxs1 >Hxs1 #IH cases (IH … (refl ??)) -IH
429       [ * #l * #l1 * #Hxs1'
430         >change_vec_commute // >change_vec_change_vec
431         #Houtc % %{(c::l)} %{l1} % 
432         [ normalize <Hxs1' %
433         | >reverse_cons >associative_append >change_vec_commute // @Houtc ]
434       | #H %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%]
435         @not_sub_list_merge
436         [ #l2 >Hxs <Hxs1 % normalize #H1 lapply (cons_injective_r ????? H1)
437           >associative_append #H2 lapply (append_l2_injective ????? (refl ??) H2)
438           #H3 lapply (cons_injective_l ????? H3) #H3 >H3 in Hcicj; * /2/
439         |#t #l2 #l3 % normalize #H1 lapply (cons_injective_r ????? H1)
440          -H1 #H1 cases (H l2 l3) #H2 @H2 @H1
441         ]
442       ]
443     | (* in match_step_true manca il caso di fallimento immediato
444          (con i due current diversi) *)
445       @daemon
446       (*
447        #_ lapply (\Pf eqx) -eqx #eqx >Hmidta_dst
448        cases (Htrue ? (refl ??) eqx) -Htrue #Htb #Hendcx #_
449        cases rs0 in Htb;
450        [ #_ %2 #l #l1 cases l
451        [ normalize cases xs
452          [ cases l1
453            [ normalize % #H destruct (H) cases eqx /2/
454            | #tmp1 #l2 normalize % #H destruct (H) ]
455          | #tmp1 #l2 normalize % #H destruct (H) ]
456        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
457          [ normalize #H1 destruct (H1)
458          | #tmp2 #l3 normalize #H1 destruct (H1) ]
459        ]
460      | #r1 #rs1 normalize in ⊢ (???(????%?)→?); #Htb >Htb in IH; #IH
461        cases (IH ls x xs end rs ? Hnotend Hend Hnotstart) 
462        [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ] -IH
463        #_ #IH cases (IH Hstart (c::ls0) r1 rs1 ?)
464        [|| >nth_change_vec // ] -IH
465        [ * #l * #l1 * #Hll1 #Hout % %{(c::l)} %{l1} % >Hll1 //
466          >reverse_cons >associative_append #cj0 #ls #Hl1 >(Hout ?? Hl1)
467          >change_vec_commute in ⊢ (??(???%??)?); // @sym_not_eq //
468        | #IH %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@IH]
469          #l1 % #H destruct (H) cases eqx /2/
470        ] *)
471     ]
472   ]
473 ]
474 qed.
475
476 definition Pre_match_m ≝ 
477   λsrc,sig,n,is_startc,is_endc.λt: Vector (tape sig) (S n).
478   ∃start,xs,end.
479   nth src (tape sig) t (niltape sig) = midtape ? [] start (xs@[end]) ∧
480   is_startc start = true ∧
481   (∀c.c ∈ (xs@[end]) = true → is_startc c = false) ∧
482   (∀c.c ∈ (start::xs) = true → is_endc c = false) ∧
483   is_endc end = true.
484   
485 lemma terminate_match_m :
486   ∀src,dst,sig,n,is_startc,is_endc,t.
487   src ≠ dst → src < S n → dst < S n → 
488   Pre_match_m src sig n is_startc is_endc t → 
489   match_m src dst sig n is_startc is_endc ↓ t.
490 #src #dst #sig #n #is_startc #is_endc #t #Hneq #Hsrc #Hdst * #start * #xs * #end
491 * * * * #Hmid_src #Hstart #Hnotstart #Hnotend #Hend
492 @(terminate_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst)) //
493 <(change_vec_same … t dst (niltape ?))
494 lapply (refl ? (nth dst (tape sig) t (niltape ?))) 
495 cases (nth dst (tape sig) t (niltape ?)) in ⊢ (???%→?);
496 [ #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
497   >Hmid_src #HR cases (HR ? (refl ??)) -HR
498   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
499   * #H @False_ind @H %
500 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
501   >Hmid_src #HR cases (HR ? (refl ??)) -HR
502   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
503   * #H @False_ind @H %
504 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
505   >Hmid_src #HR cases (HR ? (refl ??)) -HR
506   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
507   * #H @False_ind @H %
508 | #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs
509   [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
510    >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ? (refl ??)) -HR #_
511    #HR cases (HR Hstart Hnotstart)
512    cases (true_or_false (start == s)) #Hs
513    [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
514      cut (∃ci,xs1.xs@[end] = ci::xs1)
515      [ cases xs
516        [ %{end} %{[]} %
517        | #x1 #xs1 %{x1} %{(xs1@[end])} % ] ] * #ci * #xs1 #Hxs
518      >Hxs in Htrue; #Htrue
519      cases (Htrue [ ] start [ ] ? xs1 ? [ ] (refl ??) (refl ??) ?)
520      [ * #_ * #H @False_ind @H % ]
521      #c0 #Hc0 @Hnotend >(memb_single … Hc0) @memb_hd
522    | lapply (\Pf Hs) -Hs #Hs #Htrue #_
523      cases (Htrue ? (refl ??) Hs) -Htrue #Ht1 #_ %
524      #t2 whd in ⊢ (%→?); #HR cases (HR start ?)
525      [ >Ht1 >nth_change_vec // normalize in ⊢ (%→?); * #H @False_ind @H %
526      | >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
527        >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src % ]
528    ]
529   |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?);
530    >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src
531    #Htrue cases (Htrue ? (refl ??)) -Htrue #_ #Htrue
532    <(change_vec_same … t1 dst (niltape ?))
533    cases (Htrue Hstart Hnotstart) -Htrue
534    cases (true_or_false (start == s)) #Hs
535    [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
536     cut (∃ls0,xs0,ci,rs,rs0.
537       nth src ? t (niltape ?) = midtape sig [ ] start (xs0@ci::rs) ∧
538       nth dst ? t (niltape ?) = midtape sig ls0 s (xs0@rs0) ∧
539       (is_endc ci = true ∨ (is_endc ci = false ∧ (∀b,tlb.rs0 = b::tlb → ci ≠ b))))
540     [cases (comp_list ? (xs@[end]) (r0::rs0) is_endc) #xs0 * #xs1 * #xs2
541       * * * #Hxs #Hrs #Hxs0notend #Hcomp >Hrs
542       cut (∃y,ys. xs1 = y::ys)
543       [ lapply Hxs0notend lapply Hxs lapply xs0 elim xs
544         [ *
545           [ normalize #Hxs1 <Hxs1 #_ %{end} %{[]} %
546           | #z #zs normalize in ⊢ (%→?); #H destruct (H) #H
547             lapply (H ? (memb_hd …)) -H >Hend #H1 destruct (H1)
548           ]
549         | #y #ys #IH0 * 
550           [ normalize in ⊢ (%→?); #Hxs1 <Hxs1 #_ %{y} %{(ys@[end])} %
551           | #z #zs normalize in ⊢ (%→?); #H destruct (H) #Hmemb
552             @(IH0 ? e0 ?) #c #Hc @Hmemb @memb_cons // ] ] ] * #y * #ys #Hxs1
553       >Hxs1 in Hxs; #Hxs >Hmid_src >Hmid_dst >Hxs >Hrs
554       %{ls} %{xs0} %{y} %{ys} %{xs2}
555       % [ % // | @Hcomp // ] ]
556     * #ls0 * #xs0 * #ci * #rs * #rs0 * * #Hmid_src' #Hmid_dst' #Hcomp
557     <Hmid_src in Htrue; >nth_change_vec // >Hs #Htrue destruct (Hs)
558     lapply (Htrue ??????? Hmid_src' Hmid_dst' ?) -Htrue
559     [ #c0 #Hc0 @Hnotend cases (orb_true_l … Hc0) -Hc0 #Hc0
560       [ whd in ⊢ (??%?); >Hc0 %
561       | @memb_cons >Hmid_src in Hmid_src'; #Hmid_src' destruct (Hmid_src')
562         lapply e0 -e0 @(list_elim_left … rs)
563         [ #e0 destruct (e0) lapply (append_l1_injective_r ?????? e0) //
564         | #x1 #xs1 #_ >append_cons in ⊢ (???%→?);
565           <associative_append #e0 lapply (append_l1_injective_r ?????? e0) //
566           #e1 >e1 @memb_append_l1 @memb_append_l1 // ] ]
567     | * * #Hciendc cases rs0 in Hcomp;
568       [ #_ * #H @False_ind @H %
569       | #r1 #rs1 * [ >Hciendc #H destruct (H) ]
570         * #_ #Hcomp lapply (Hcomp ?? (refl ??)) -Hcomp #Hcomp #_ #Htrue
571         cases (Htrue ?? (refl ??) Hcomp) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
572         [ >nth_change_vec_neq [|@sym_not_eq //]
573           >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
574         | >nth_change_vec // >Hmid_dst % ] ] ]
575   | >Hmid_dst >nth_change_vec // lapply (\Pf Hs) -Hs #Hs #Htrue #_
576     cases (Htrue ? (refl ??) Hs) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
577     [ >nth_change_vec_neq [|@sym_not_eq //]
578       >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
579     | >nth_change_vec // ] ] ] ]
580 qed.