]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/match.ma
match nearing completion
[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
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 lemma comp_list: ∀S:DeqSet. ∀l1,l2:list S.∀is_endc. ∃l,tl1,tl2. 
75   l1 = l@tl1 ∧ l2 = l@tl2 ∧ (∀c.c ∈ l = true → is_endc c = false) ∧
76   ∀a,tla. tl1 = a::tla → 
77   is_endc a = true ∨ (is_endc a = false ∧∀b,tlb.tl2 = b::tlb → a≠b).
78 #S #l1 #l2 #is_endc elim l1 in l2;
79 [ #l2 %{[ ]} %{[ ]} %{l2} normalize %
80   [ % [ % // | #c #H destruct (H) ] | #a #tla #H destruct (H) ]
81 | #x #l3 #IH cases (true_or_false (is_endc x)) #Hendcx
82   [ #l %{[ ]} %{(x::l3)} %{l} normalize
83     % [ % [ % // | #c #H destruct (H) ] | #a #tla #H destruct (H) >Hendcx % % ]
84   | *
85     [ %{[ ]} %{(x::l3)} %{[ ]} normalize %
86       [ % [ % // | #c #H destruct (H) ]
87       | #a #tla #H destruct (H) cases (is_endc a)
88         [ % % | %2 % // #b #tlb #H destruct (H) ]
89       ]
90     | #y #l4 cases (true_or_false (x==y)) #Hxy
91       [ lapply (\P Hxy) -Hxy #Hxy destruct (Hxy)  
92         cases (IH l4) -IH #l * #tl1 * #tl2 * * * #Hl3 #Hl4 #Hl #IH
93         %{(y::l)} %{tl1} %{tl2} normalize
94         % [ % [ % // 
95               | #c cases (true_or_false (c==y)) #Hcy >Hcy normalize
96                 [ >(\P Hcy) //
97                 | @Hl ]
98               ]
99           | #a #tla #Htl1 @(IH … Htl1) ]
100       | %{[ ]} %{(x::l3)} %{(y::l4)} normalize %
101         [ % [ % // | #c #H destruct (H) ]
102         | #a #tla #H destruct (H) cases (is_endc a)
103           [ % % | %2 % // #b #tlb #H destruct (H) @(\Pf Hxy) ]
104         ]
105       ]
106     ]
107   ]
108 ]
109 qed.        
110   
111 axiom daemon : ∀X:Prop.X.
112
113 definition match_test ≝ λsrc,dst.λsig:DeqSet.λn,is_endc.λv:Vector ? n.
114   match (nth src (option sig) v (None ?)) with 
115   [ None ⇒  false 
116   | Some x ⇒  notb ((is_endc x) ∨ (nth dst (DeqOption sig) v (None ?) == None ?))]. 
117
118 definition match_step ≝ λsrc,dst,sig,n,is_startc,is_endc.
119   compare src dst sig n is_endc ·
120      (ifTM ?? (partest sig n (match_test src dst sig ? is_endc))
121       (single_finalTM ??
122         (parmove src dst sig n L is_startc · (inject_TM ? (move_r ?) n dst)))
123       (nop …)
124       partest1).
125       
126 definition R_match_step_false ≝ 
127   λsrc,dst,sig,n,is_endc.λint,outt: Vector (tape sig) (S n).
128   ∀ls,x,xs,end,rs.
129   nth src ? int (niltape ?) = midtape sig ls x (xs@end::rs) →
130   (∀c0. memb ? c0 (x::xs) = true → is_endc c0 = false) → is_endc end = true →
131    ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
132     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
133       xs = rs0@xs0 ∧
134       current sig (nth dst (tape sig) outt (niltape sig)) = None ?) ∨
135     (∃ls0,rs0. 
136      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
137      ∀rsj,c. 
138      rs0 = c::rsj →
139      outt = change_vec ??
140             (change_vec ?? int (midtape sig (reverse ? xs@x::ls) end rs) src)
141             (midtape sig (reverse ? xs@x::ls0) c rsj) dst).
142
143 definition R_match_step_true ≝ 
144   λsrc,dst,sig,n,is_startc,is_endc.λint,outt: Vector (tape sig) (S n).
145   ∀s.current sig (nth src (tape sig) int (niltape sig)) = Some ? s → 
146   current sig (nth dst (tape sig) int (niltape sig)) ≠ None ? ∧
147   (is_startc s = true → 
148    (∀c.c ∈ right ? (nth src (tape sig) int (niltape sig)) = true → is_startc c = false) →
149    (∀s1.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s1 → s ≠ s1 →  
150     outt = change_vec ?? int 
151           (tape_move … (nth dst ? int (niltape ?)) (Some ? 〈s1,R〉)) dst ∧ is_endc s = false) ∧  
152    (∀ls,x,xs,ci,cj,rs,ls0,rs0. 
153      nth src ? int (niltape ?) = midtape sig ls x (xs@ci::rs) →
154      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@cj::rs0) →
155      (∀c0. memb ? c0 (x::xs) = true → is_endc c0 = false) → 
156       ci ≠ cj →
157       (outt = change_vec ?? int 
158           (tape_move … (nth dst ? int (niltape ?)) (Some ? 〈x,R〉)) dst ∧ is_endc ci = false))). 
159 (*    ∧
160     (rs0 = [ ] →
161      outt = change_vec ??
162            (change_vec ?? int (midtape sig (reverse ? xs@x::ls) ci rs) src)
163            (mk_tape sig (reverse ? xs@x::ls0) (None ?) [ ]) dst)). *)
164            
165 lemma sem_match_step :
166   ∀src,dst,sig,n,is_startc,is_endc.src ≠ dst → src < S n → dst < S n → 
167   match_step src dst sig n is_startc is_endc ⊨ 
168     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
169       R_match_step_true src dst sig n is_startc is_endc, 
170       R_match_step_false src dst sig n is_endc ].
171 #src #dst #sig #n #is_startc #is_endc #Hneq #Hsrc #Hdst 
172 @(acc_sem_seq_app sig n … (sem_compare src dst sig n is_endc Hneq Hsrc Hdst)
173     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ? is_endc))
174       (sem_seq … 
175         (sem_parmoveL ???? is_startc Hneq Hsrc Hdst) 
176         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
177       (sem_nop …)))
178 [#ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd >Htd -Htd
179  * #te * #Hte #Htb whd 
180  #s #Hcurta_src % 
181  [ lapply (refl ? (current ? (nth dst ? ta (niltape ?)))) 
182    cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
183    [| #c #_ % #Hfalse destruct (Hfalse) ]
184    #Hcurta_dst >Hcomp1 in Htest; [| %2 %2 //]
185    whd in ⊢ (??%?→?); change with (current ? (niltape ?)) in match (None ?);
186     <nth_vec_map >Hcurta_src whd in ⊢ (??%?→?); <nth_vec_map
187     >Hcurta_dst cases (is_endc s) normalize in ⊢ (%→?); #H destruct (H)
188  | #Hstart #Hnotstart %
189    [ #s1 #Hcurta_dst #Hneqss1 -Hcomp2
190      cut (tc = ta) 
191      [@Hcomp1 %2 %1 %1 >Hcurta_src >Hcurta_dst @(not_to_not … Hneqss1) #H destruct (H) //] 
192      #H destruct (H) -Hcomp1 cases Hte #_ -Hte #Hte
193      cut (te = ta) [@Hte %1 %1 %{s} % //] -Hte #H destruct (H) %
194      [cases Htb * #_ #Hmove #Hmove1 @(eq_vec … (niltape … ))
195       #i #Hi cases (decidable_eq_nat i dst) #Hidst
196        [ >Hidst >nth_change_vec // cases (current_to_midtape … Hcurta_dst)
197          #ls * #rs #Hta_mid >(Hmove … Hta_mid) >Hta_mid cases rs //
198        | >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @Hmove1 @sym_not_eq // ]
199      | whd in Htest:(??%?); >(nth_vec_map ?? (current sig)) in Hcurta_src; #Hcurta_src
200        >Hcurta_src in Htest; whd in ⊢ (??%?→?);
201        cases (is_endc s) // whd in ⊢ (??%?→?); #H @sym_eq // 
202      ]
203    |#ls #x #xs #ci #cj #rs #ls0 #rs00 #Htasrc_mid #Htadst_mid #Hnotendc #Hcicj 
204     cases (Hcomp2 … Htasrc_mid Htadst_mid Hnotendc) [ * #H destruct (H) ]
205     * #cj' * #rs0' * #Hcjrs0 destruct (Hcjrs0) -Hcomp2 #Hcomp2
206     lapply (Hcomp2 (or_intror ?? Hcicj)) -Hcomp2 #Htc %
207     [ cases Hte -Hte #Hte #_ whd in Hte;
208       >Htasrc_mid in Hcurta_src; whd in ⊢ (??%?→?); #H destruct (H) 
209       lapply (Hte ls ci (reverse ? xs) rs s ??? ls0 cj' (reverse ? xs) s rs0' (refl ??) ?) //
210       [ >Htc >nth_change_vec //
211       | #c0 #Hc0 @(Hnotstart c0) >Htasrc_mid cases (orb_true_l … Hc0) -Hc0 #Hc0
212         [@memb_append_l2 >(\P Hc0) @memb_hd
213         |@memb_append_l1 <(reverse_reverse …xs) @memb_reverse //
214         ]
215       | >Htc >change_vec_commute // >nth_change_vec // ] -Hte
216       >Htc >change_vec_commute // >change_vec_change_vec 
217       >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
218       >Hte in Htb; * * #_ >reverse_reverse #Htbdst1 #Htbdst2 -Hte @(eq_vec … (niltape ?))
219       #i #Hi cases (decidable_eq_nat i dst) #Hidst
220       [ >Hidst >nth_change_vec // >(Htbdst1 ls0 s (xs@cj'::rs0'))
221         [| >nth_change_vec // ]
222         >Htadst_mid cases xs //
223       | >nth_change_vec_neq [|@sym_not_eq // ]
224         <Htbdst2 [| @sym_not_eq // ] >nth_change_vec_neq [| @sym_not_eq // ]
225         <Htasrc_mid >change_vec_same % ]
226     | >Hcurta_src in Htest; whd in ⊢(??%?→?);
227       >Htc >change_vec_commute //
228       change with (current ? (niltape ?)) in match (None ?);
229       <nth_vec_map >nth_change_vec // whd in ⊢ (??%?→?);
230       cases (is_endc ci) whd in ⊢ (??%?→?); #H destruct (H) % 
231     ]
232    ]
233   ]
234 |#intape #outtape #ta * #Hcomp1 #Hcomp2 * #tb * * #Hc #Htb 
235  whd in ⊢ (%→?); #Hout >Hout >Htb whd
236  #ls #c_src #xs #end #rs #Hmid_src #Hnotend #Hend
237  lapply (current_to_midtape sig (nth dst ? intape (niltape ?)))
238  cases (current … (nth dst ? intape (niltape ?))) in Hcomp1;
239   [#Hcomp1 #_ %1 % % [% | @Hcomp1 %2 %2 % ]
240   |#c_dst cases (true_or_false (c_src == c_dst)) #Hceq
241     [#_ #Hmid_dst cases (Hmid_dst c_dst (refl …)) -Hmid_dst
242      #ls_dst * #rs_dst #Hmid_dst
243      cases (comp_list … (xs@end::rs) rs_dst is_endc) #xs1 * #rsi * #rsj * * * 
244      #Hrs_src #Hrs_dst #Hnotendxs1 #Hneq >Hrs_dst in Hmid_dst; #Hmid_dst
245      cut (∃r1,rs1.rsi = r1::rs1) [@daemon] * #r1 * #rs1 #Hrs1 >Hrs1 in Hrs_src;
246      #Hrs_src >Hrs_src in Hmid_src; #Hmid_src <(\P Hceq) in Hmid_dst; #Hmid_dst
247      lapply (Hcomp2 ??????? Hmid_src Hmid_dst ?) 
248      [ #c0 #Hc0 cases (orb_true_l … Hc0) -Hc0 #Hc0
249        [ >(\P Hc0) @Hnotend @memb_hd | @Hnotendxs1 //] ] 
250      *
251      [ * #Hrsj >Hrsj #Hta % %2 >Hta >nth_change_vec //
252        %{ls_dst} %{xs1} cut (∃xs0.xs = xs1@xs0)
253        [lapply Hnotendxs1 -Hnotendxs1 lapply Hrs_src lapply xs elim xs1
254          [ #l #_ #_ %{l} %
255          | #x2 #xs2 #IH * 
256            [ whd in ⊢ (??%%→?); #H destruct (H) #Hnotendxs2
257              >Hnotendxs2 in Hend; [ #H destruct (H) |@memb_hd ]
258            | #x2' #xs2' whd in ⊢ (??%%→?); #H destruct (H)
259              #Hnotendxs2 cases (IH xs2' e0 ?)
260              [ #xs0 #Hxs2 %{xs0} @eq_f //
261              |#c #Hc @Hnotendxs2 @memb_cons // ]
262            ]
263          ] 
264        ] * #xs0 #Hxs0 %{xs0} % [ %
265        [ >Hmid_dst >Hrsj >append_nil %
266        | @Hxs0 ]
267        | cases (reverse ? xs1) // ]
268      | * #cj * #rs2 * #Hrsj #Hta lapply (Hta ?)
269        [ cases (Hneq ?? Hrs1) /2/ * #_ #Hr1 %2 @(Hr1 ?? Hrsj) ] -Hta #Hta
270        %2 >Hta in Hc; whd in ⊢ (??%?→?);
271        change with (current ? (niltape ?)) in match (None ?);
272        <nth_vec_map >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
273        whd in ⊢ (??%?→?); #Hc cut (is_endc r1 = true)
274        [ cases (is_endc r1) in Hc; whd in ⊢ (??%?→?); //
275          change with (current ? (niltape ?)) in match (None ?);
276          <nth_vec_map >nth_change_vec // normalize #H destruct (H) ]
277        #Hendr1 cut (xs = xs1)
278        [ lapply Hnotendxs1 lapply Hnotend lapply Hrs_src lapply xs1
279          -Hnotendxs1 -Hnotend -Hrs_src -xs1 elim xs
280          [ * normalize in ⊢ (%→?); //
281            #x2 #xs2 normalize in ⊢ (%→?); #Heq destruct (Heq) #_ #Hnotendxs1
282            lapply (Hnotendxs1 ? (memb_hd …)) >Hend #H destruct (H)
283          | #x2 #xs2 #IH *
284            [ normalize in ⊢ (%→?); #Heq destruct (Heq) #Hnotendc
285              >Hnotendc in Hendr1; [| @memb_cons @memb_hd ]
286              normalize in ⊢ (%→?); #H destruct (H)
287            | #x3 #xs3 normalize in ⊢ (%→?); #Heq destruct (Heq)
288              #Hnotendc #Hnotendcxs1 @eq_f @IH
289              [ @(cons_injective_r … Heq)
290              | #c0 #Hc0 @Hnotendc cases (orb_true_l … Hc0) -Hc0 #Hc0
291                [ >(\P Hc0) @memb_hd
292                | @memb_cons @memb_cons // ]
293              | #c #Hc @Hnotendcxs1 @memb_cons // ]
294            ]
295          ]
296        | #Hxsxs1 destruct (Hxsxs1) >Hmid_dst %{ls_dst} %{rsj} % //
297          #rsj0 #c >Hrsj #Hrsj0 destruct (Hrsj0) 
298          lapply (append_l2_injective … Hrs_src) // #Hrs' destruct (Hrs') %
299        ]
300      ]
301     |#Hcomp1 #Hsrc cases (Hsrc ? (refl ??)) -Hsrc #ls0 * #rs0 #Hdst 
302      @False_ind lapply (Hcomp1 ?) [%2 %1 %1 >Hmid_src normalize
303      @(not_to_not ??? (\Pf Hceq)) #H destruct //] #Hintape >Hintape in Hc;
304      whd in ⊢(??%?→?); >Hmid_src  
305      change with (current ? (niltape ?)) in match (None ?);
306      <nth_vec_map >Hmid_src whd in ⊢ (??%?→?);
307      >(Hnotend c_src) [|@memb_hd]
308      change with (current ? (niltape ?)) in match (None ?);
309      <nth_vec_map >Hmid_src whd in ⊢ (??%?→?); >Hdst normalize #H destruct (H)
310    ]
311   ]
312 ]
313 qed.
314
315 definition match_m ≝ λsrc,dst,sig,n,is_startc,is_endc.
316   whileTM … (match_step src dst sig n is_startc is_endc) 
317     (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
318
319 definition R_match_m ≝ 
320   λsrc,dst,sig,n,is_startc,is_endc.λint,outt: Vector (tape sig) (S n).
321 (*  (current sig (nth dst (tape sig) int (niltape sig)) = None ? → outt = int) ∧ *)
322   ∀ls,x,xs,end,rs.
323   nth src ? int (niltape ?) = midtape sig ls x (xs@end::rs) →
324   (∀c0. memb ? c0 (x::xs) = true → is_endc c0 = false) → is_endc end = true →
325   (∀c0. memb ? c0 (xs@end::rs) = true → is_startc c0 = false) → 
326   (current sig (nth dst (tape sig) int (niltape sig)) = None ? → outt = int) ∧
327   (is_startc x = true →
328    (∀ls0,x0,rs0.
329     nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
330     (∃l,l1.x0::rs0 = l@x::xs@l1 ∧
331      ∀cj,l2.l1=cj::l2 →
332      outt = change_vec ?? 
333             (change_vec ?? int (midtape sig (reverse ? xs@x::ls) end rs) src)
334             (midtape sig ((reverse ? (l@x::xs))@ls0) cj l2) dst) ∨
335     ∀l,l1.x0::rs0 ≠ l@x::xs@l1)).
336
337 (*
338 definition R_match_m ≝ 
339   λi,j,sig,n,is_startc,is_endc.λint,outt: Vector (tape sig) (S n).
340   (((∃x.current ? (nth i ? int (niltape ?)) = Some ? x ∧ is_endc x = true) ∨
341     current ? (nth i ? int (niltape ?)) = None ? ∨
342     current ? (nth j ? int (niltape ?)) = None ?) → outt = int) ∧
343   (∀ls,x,xs,ci,rs,ls0,x0,rs0.
344     (∀x. is_startc x ≠ is_endc x) → 
345     is_startc x = true → is_endc ci = true → 
346     (∀z. memb ? z (x::xs) = true → is_endc x = false) →
347     nth i ? int (niltape ?) = midtape sig ls x (xs@ci::rs) →
348     nth j ? int (niltape ?) = midtape sig ls0 x0 rs0 →
349     (∃l,l1.x0::rs0 = l@x::xs@l1 ∧
350      ∀cj,l2.l1=cj::l2 →
351      outt = change_vec ?? 
352             (change_vec ?? int (midtape sig (reverse ? xs@x::ls) ci rs) i)
353             (midtape sig ((reverse ? (l@x::xs))@ls0) cj l2) j) ∨
354     ∀l,l1.x0::rs0 ≠ l@x::xs@l1).
355 *)
356
357 (*
358 axiom sub_list_dec: ∀A.∀l,ls:list A. 
359   ∃l1,l2. l = l1@ls@l2 ∨ ∀l1,l2. l ≠ l1@ls@l2.
360 *)
361
362 lemma not_sub_list_merge : 
363   ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
364 #T #a #b #H1 #H2 #l elim l normalize //
365 qed.
366
367 lemma not_sub_list_merge_2 : 
368   ∀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.
369 #T #a #b #t #H1 #H2 #l elim l //
370 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
371 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
372 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
373 qed.
374
375
376 lemma wsem_match_m : ∀src,dst,sig,n,is_startc,is_endc.
377 src ≠ dst → src < S n → dst < S n → 
378   match_m src dst sig n is_startc is_endc ⊫ R_match_m src dst sig n is_startc is_endc.
379 #src #dst #sig #n #is_startc #is_endc #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
380 lapply (sem_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst) … Hloop) //
381 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
382 [ #tc #Hfalse #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart
383   cases (Hfalse … Hmid_src Hnotend Hend) -Hfalse 
384   [(* current dest = None *) *
385     [ * #Hcur_dst #Houtc %
386       [#_ >Houtc //
387       |#Hstart #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst; 
388        normalize in ⊢ (%→?); #H destruct (H)
389       ]
390     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
391       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
392       | #Hstart #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
393         >Hrs0 cases xs0
394         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
395           #cj #ls2 #H destruct (H)
396         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
397           >length_append whd in ⊢ (??%(??%)→?); >length_append
398           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
399           #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
400           >associative_plus >associative_plus 
401           #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
402           #e2 destruct (e2)
403         ]
404       ]
405     ]
406   |* #ls0 * #rs0 * #Hmid_dst #HFalse %
407     [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
408     | #Hstart #ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
409      %1 %{[ ]} %{rs0} % [%] #cj #l2 #Hnotnil 
410      >reverse_cons >associative_append @(HFalse ?? Hnotnil)
411     ]
412   ]
413 |#ta #tb #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
414  #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart 
415  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
416  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
417   [#Hmid_dst % 
418     [#_ whd in Htrue; >Hmid_src in Htrue; #Htrue
419      cases (Htrue x (refl … )) -Htrue * #Htaneq #_
420      @False_ind >Hmid_dst in Htaneq; /2/
421     |#Hstart #ls0 #x0 #rs0 #Hmid_dst2 >Hmid_dst2 in Hmid_dst; normalize in ⊢ (%→?); 
422      #H destruct (H)
423     ]
424   | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
425     #Hstart #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcurta_dst; normalize in ⊢ (%→?);
426     #H destruct (H) whd in Htrue; >Hmid_src in Htrue; #Htrue
427     cases (Htrue x (refl …)) -Htrue #_ #Htrue cases (Htrue Hstart Hnotstart) -Htrue
428     cases (true_or_false (x==c)) #eqx
429     [ lapply (\P eqx) -eqx #eqx destruct (eqx)
430       #_ #Htrue cases (comp_list ? (xs@end::rs) rs0 is_endc)
431       #x1 * #tl1 * #tl2 * * * #Hxs #Hrs0 #Hnotendx1
432       cases tl1 in Hxs; 
433       [>append_nil #Hx1 <Hx1 in Hnotendx1; #Hnotendx1
434        lapply (Hnotendx1 end ?) [ @memb_append_l2 @memb_hd ]
435        >Hend #H destruct (H) ]
436       #ci -tl1 #tl1 #Hxs #H cases (H … (refl … ))
437       [ #Hendci % cases (IH ????? Hmid_src Hnotend Hend Hnotstart)
438       (* this is absurd, since Htrue conlcudes is_endc ci =false *)
439        (* no, è più complicato
440        #Hend_ci lapply (Htrue ls c xi
441        *)
442         @daemon (* lapply(Htrue … (refl …)) -Htrue *)
443       |cases tl2 in Hrs0;
444         [ >append_nil #Hrs0 destruct (Hrs0) * #Hcifalse#_ %2
445           cut (∃l.xs = x1@ci::l) 
446           [lapply Hxs lapply Hnotendx1 lapply Hnotend lapply xs 
447            -Hxs -xs -Hnotendx1 elim x1
448             [ *
449               [ #_ #_ normalize #H1 destruct (H1) >Hend in Hcifalse;
450                 #H1 destruct (H1)
451               | #x2 #xs2 #_ #_ normalize #H >(cons_injective_l ????? H) %{xs2} % ]
452             | #x2 #xs2 #IHin *
453               [ #_ #Hnotendxs2 normalize #H destruct (H) 
454                 >(Hnotendxs2 ? (memb_hd …)) in Hend; #H destruct (H)
455               | #x3 #xs3 #Hnotendxs3 #Hnotendxs2 normalize #H destruct (H)
456                 cases (IHin ??? e0)
457                 [ #xs4 #Hxs4 >Hxs4 %{xs4} %
458                 | #c0 #Hc0 cases (orb_true_l … Hc0) -Hc0 #Hc0
459                   [ >(\P Hc0) @Hnotendxs3 @memb_hd
460                   | @Hnotendxs3 @memb_cons @memb_cons @Hc0 ]
461                 | #c0 #Hc0 @Hnotendxs2 @memb_cons @Hc0 ]
462               ]
463             ]
464           ] * #l #Hxs' >Hxs'
465           #l0 #l1 % #H lapply (eq_f ?? (length ?) ?? H) -H
466           >length_append normalize >length_append >length_append
467           normalize >commutative_plus normalize #H destruct (H) -H
468           >associative_plus in e0; >associative_plus
469           >(plus_n_O (|x1|)) in ⊢(??%?→?); #H lapply (injective_plus_r … H)
470           -H normalize #H destruct (H)
471         | #cj #tl2' #Hrs0 * #Hcifalse #Hcomp lapply (Htrue ls c x1 ci cj tl1 ls0 tl2' ????)
472           [ @(Hcomp ?? (refl ??))
473           | #c0 #Hc0 cases (orb_true_l … Hc0) #Hc0
474             [ @Hnotend >(\P Hc0) @memb_hd
475             | @Hnotendx1 // ]
476           | >Hmid_dst >Hrs0 %
477           | >Hxs %
478           | * #Htb >Htb #Hendci >Hrs0 >Hxs
479             cases (IH ls c xs end rs ? Hnotend Hend Hnotstart) -IH 
480             [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ]
481             #_ #IH lapply Hxs lapply Hnotendx1 -Hxs -Hnotendx1 cases x1 in Hrs0;
482             [ #Hrs0 #_ whd in ⊢ (???%→?); #Hxs
483               cases (IH Hstart (c::ls0) cj tl2' ?)
484               [ -IH * #l * #l1 * #Hll1 #IH % %{(c::l)} %{l1} 
485                 % [ @eq_f @Hll1 ]
486                 #cj0 #l2 #Hcj0 >(IH … Hcj0) >Htb
487                 >change_vec_commute // >change_vec_change_vec
488                 >change_vec_commute [|@sym_not_eq // ] @eq_f3 //
489                 >reverse_cons >associative_append %
490               | #IH %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%]
491                 @not_sub_list_merge
492                 [ #l2 cut (∃xs'.xs = ci::xs')
493                   [ cases xs in Hxs;
494                     [ normalize #H destruct (H) >Hend in Hendci; #H destruct (H)
495                     | #ci' #xs' normalize #H lapply (cons_injective_l ????? H)
496                       #H1 >H1 %{xs'} % ]
497                   ]
498                   * #xs' #Hxs' >Hxs' normalize % #H destruct (H)
499                   lapply (Hcomp … (refl ??)) * /2/
500                 |#t #l2 #l3 % normalize #H lapply (cons_injective_r ????? H)
501                  -H #H >H in IH; #IH cases (IH l2 l3) -IH #IH @IH % ] 
502               | >Htb >nth_change_vec // >Hmid_dst >Hrs0 % ]
503             | #x2 #xs2 normalize in ⊢ (%→?); #Hrs0 #Hnotendxs2 normalize in ⊢ (%→?);
504               #Hxs cases (IH Hstart (c::ls0) x2 (xs2@cj::tl2') ?)
505               [ -IH * #l * #l1 * #Hll1 #IH % %{(c::l)} %{l1} 
506                 % [ @eq_f @Hll1 ]
507                 #cj0 #l2 #Hcj0 >(IH … Hcj0) >Htb
508                 >change_vec_commute // >change_vec_change_vec
509                 >change_vec_commute [|@sym_not_eq // ] @eq_f3 //
510                 >reverse_cons >associative_append %
511               | -IH #IH %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%] 
512                 @not_sub_list_merge_2 [| @IH]
513                 cut (∃l2.xs = (x2::xs2)@ci::l2)
514                 [lapply Hnotendxs2
515                  lapply Hnotend -Hnotend lapply Hxs
516                  >(?:x2::xs2@ci::tl1 = (x2::xs2)@ci::tl1) [|%]
517                  lapply (x2::xs2) elim xs
518                   [ *
519                     [ normalize in ⊢ (%→?); #H1 destruct (H1) 
520                       >Hendci in Hend; #Hend destruct (Hend)
521                     | #x3 #xs3 normalize in ⊢ (%→?); #H1 destruct (H1)
522                       #_ #Hnotendx3 >(Hnotendx3 ? (memb_hd …)) in Hend;
523                       #Hend destruct (Hend)
524                     ]
525                   | #x3 #xs3 #IHin *
526                     [ normalize in ⊢ (%→?); #Hxs3 destruct (Hxs3) #_ #_
527                       %{xs3} %
528                     | #x4 #xs4 normalize in ⊢ (%→?); #Hxs3xs4 #Hnotend
529                       #Hnotendxs4 destruct (Hxs3xs4) cases (IHin ? e0 ??)
530                       [ #l0 #Hxs3 >Hxs3 %{l0} %
531                       | #c0 #Hc0 @Hnotend cases (orb_true_l … Hc0) -Hc0 #Hc0
532                         [ >(\P Hc0) @memb_hd
533                         | @memb_cons @memb_cons @Hc0 ]
534                       | #c0 #Hc0 @Hnotendxs4 @memb_cons //
535                       ]
536                     ]
537                   ]
538                 ] * #l2 #Hxs'
539                 >Hxs' #l3 normalize >associative_append normalize % #H
540                 destruct (H) lapply (append_l2_injective ?????? e1) //
541                 #H1 destruct (H1) cases (Hcomp ?? (refl ??)) /2/
542               | >Htb >nth_change_vec // >Hmid_dst >Hrs0 % ]
543             ]
544           ]
545         ]
546       ]
547     |lapply (\Pf eqx) -eqx #eqx >Hmid_dst #Htrue 
548      cases (Htrue ? (refl ??) eqx) -Htrue #Htb #Hendcx #_
549      cases rs0 in Htb;
550      [ #_ %2 #l #l1 cases l
551        [ normalize cases xs
552          [ cases l1
553            [ normalize % #H destruct (H) cases eqx /2/
554            | #tmp1 #l2 normalize % #H destruct (H) ]
555          | #tmp1 #l2 normalize % #H destruct (H) ]
556        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
557          [ normalize #H1 destruct (H1)
558          | #tmp2 #l3 normalize #H1 destruct (H1) ]
559        ]
560      | #r1 #rs1 normalize in ⊢ (???(????%?)→?); #Htb >Htb in IH; #IH
561        cases (IH ls x xs end rs ? Hnotend Hend Hnotstart) 
562        [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ] -IH
563        #_ #IH cases (IH Hstart (c::ls0) r1 rs1 ?)
564        [|| >nth_change_vec // ] -IH
565        [ * #l * #l1 * #Hll1 #Hout % %{(c::l)} %{l1} % >Hll1 //
566          >reverse_cons >associative_append #cj0 #ls #Hl1 >(Hout ?? Hl1)
567          >change_vec_commute in ⊢ (??(???%??)?); // @sym_not_eq //
568        | #IH %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@IH]
569          #l1 % #H destruct (H) cases eqx /2/
570        ] 
571     ]
572   ]
573 ]
574 qed.
575