]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/universal/copy.ma
New notation for congruence
[helm.git] / matita / matita / lib / turing / universal / copy.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic   
3     ||A||  Library of Mathematics, developed at the Computer Science 
4     ||T||  Department of the University of Bologna, Italy.           
5     ||I||                                                            
6     ||T||  
7     ||A||  
8     \   /  This file is distributed under the terms of the       
9      \ /   GNU General Public License Version 2   
10       V_____________________________________________________________*)
11
12
13 include "turing/universal/marks.ma".
14
15 definition copy_step_subcase ≝
16   λalpha,c,elseM.ifTM ? (test_char ? (λx.x == 〈c,true〉))
17     (seq (FinProd alpha FinBool) (adv_mark_r …)
18       (seq ? (move_l …)
19         (seq ? (adv_to_mark_l … (is_marked alpha))
20           (seq ? (write ? 〈c,false〉)
21             (seq ? (move_r …)
22               (seq ? (mark …)
23                 (seq ? (move_r …) (adv_to_mark_r … (is_marked alpha)))))))))
24     elseM tc_true.
25
26 definition R_copy_step_subcase ≝ 
27   λalpha,c,RelseM,t1,t2.
28     ∀a,l1,x0,a0,l2,x,l3.
29     t1 = midtape (FinProd … alpha FinBool) (l1@〈a0,false〉::〈x0,true〉::l2) 
30          〈x,true〉 (〈a,false〉::l3) → 
31     (∀c.memb ? c l1 = true → is_marked ? c = false) →          
32     (x = c ∧ t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x,false〉::l2) 〈a,true〉 l3) ∨
33     (x ≠ c ∧ RelseM t1 t2).
34     
35 lemma sem_copy_step_subcase : 
36   ∀alpha,c,elseM,RelseM. Realize ? elseM RelseM → 
37   Realize ? (copy_step_subcase alpha c elseM) (R_copy_step_subcase alpha c RelseM).
38 #alpha #c #elseM #RelseM #HelseM #intape
39 cases (sem_if ? (test_char ? (λx. x == 〈c,true〉)) ?????? tc_true (sem_test_char ? (λx.x == 〈c,true〉))
40         (sem_seq ????? (sem_adv_mark_r alpha)
41           (sem_seq ????? (sem_move_l …)
42             (sem_seq ????? (sem_adv_to_mark_l … (is_marked alpha))
43               (sem_seq ????? (sem_write ? 〈c,false〉)
44                 (sem_seq ????? (sem_move_r …)
45                   (sem_seq ????? (sem_mark …)
46                     (sem_seq ????? (sem_move_r …) (sem_adv_to_mark_r … (is_marked alpha)))))))))
47         HelseM intape)
48 #k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop
49 #a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR
50 [ * #ta * whd in ⊢ (%→?); >Hintape * * #c0 * whd in ⊢ (??%?→?); #Hx #Hc #Hta
51   * #tb * whd in ⊢ (%→?); * #Htb cases (Htb (l1@〈a0,false〉::〈x0,true〉::l2) x) -Htb
52   #Htb lapply (Htb … Hta) -Htb #Htb #_ #_
53   * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc #Htc
54   * #td * whd in ⊢ (%→?); * #_ #Htd cases (Htd … Htc) -Htd #_ #Htd cases (Htd (refl ??))
55   -Htd #Htd lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ???) //
56   [#x1 #Hx1 cases (memb_append … Hx1) [ @Hl1marks | #Hsingle >(memb_single … Hsingle) % ]
57   |whd in ⊢ (??%?); // ]
58   -Htd #Htd #_
59   * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd -Htc #Hte
60   * #tf * whd in ⊢ (%→?); * #_ #Htf lapply (Htf … Hte) -Hte -Htf >reverse_append #Htf
61   * #tg * whd in ⊢ (%→?); * #Htg #_ lapply (Htg … Htf) -Htf -Htg >reverse_single #Htg
62   * #th * whd in ⊢ (%→?); * #_ #Hth lapply (Hth … Htg) -Htg -Hth
63   generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1)
64   [ #Hl1marks #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
65     [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
66     * * #_ #Houtc #_ lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc
67     #Houtc % >(\P Hc) in Hx; #Hx destruct (Hx) % // @Houtc
68   | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single
69     #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
70     [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ]
71     * * #_ #Houtc lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc
72     [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1   
73       [@Hl1marks @memb_append_l1 <(reverse_reverse … l1) @memb_reverse @Hx1 
74       |>(memb_single … Hx1) % ]
75     | normalize >associative_append % ] 
76     #Houtc #_ % destruct (Hx) lapply (\P Hc) -Hc #Hc destruct (Hc) % //
77     >Houtc >reverse_append >reverse_reverse >associative_append >associative_append % ]
78 | * #ta * whd in ⊢ (%→?); >Hintape * #Hxc #Hta #Helse %2 %
79   [| <Hta @Helse ]
80   % #Hxc0 >Hxc0 in Hxc; #Hxc lapply (Hxc 〈c,true〉 (refl …)) #Hfalse
81   cases (\Pf Hfalse) #Hfalse0 @Hfalse0 %
82 ]
83 qed.
84     
85 (*
86 if current = 0,tt
87    then advance_mark_r;
88         move_l;
89         advance_to_mark_l;
90         write(0,ff)
91         move_r;
92         mark;
93         move_r;
94         advance_to_mark_r;
95 else if current = 1,tt
96    then advance_mark_r;
97         move_l;
98         advance_to_mark_l;
99         write(1,ff)
100         move_r;
101         mark;
102         move_r;
103         advance_to_mark_r;
104 else if current = null 
105    then advance_mark_r;
106         move_l;
107         advance_to_mark_l
108         adv_mark_r;
109         move_r;
110         advance_to_mark_r
111 *)
112
113 definition nocopy_subcase ≝
114   ifTM STape (test_char ? (λx:STape.x == 〈null,true〉))
115     (seq ? (adv_mark_r …)
116       (seq ? (move_l …)
117         (seq ? (adv_to_mark_l … (is_marked ?))
118           (seq ? (adv_mark_r …)
119             (seq ? (move_r …) (adv_to_mark_r … (is_marked ?)))))))
120     (nop ?) tc_true.
121
122 definition R_nocopy_subcase ≝ 
123   λt1,t2.
124     ∀a,l1,x0,a0,l2,x,l3.
125     t1 = midtape STape (l1@〈a0,false〉::〈x0,true〉::l2) 
126          〈x,true〉 (〈a,false〉::l3) → 
127     (∀c.memb ? c l1 = true → is_marked ? c = false) →          
128     (x = null ∧
129      t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3) ∨
130     (x ≠ null ∧ t2 = t1).
131     
132 lemma sem_nocopy_subcase : Realize ? nocopy_subcase R_nocopy_subcase.
133 #intape
134 cases (sem_if ? (test_char ? (λx:STape.x == 〈null,true〉)) ?????? tc_true
135         (sem_test_char ? (λx:STape.x == 〈null,true〉))
136           (sem_seq … (sem_adv_mark_r …)
137             (sem_seq … (sem_move_l …)
138               (sem_seq … (sem_adv_to_mark_l … (is_marked ?))
139                 (sem_seq … (sem_adv_mark_r …)
140                   (sem_seq … (sem_move_r …) 
141                     (sem_adv_to_mark_r … (is_marked ?))))))) (sem_nop ?) intape)
142 #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc)  % [@Hloop] -Hloop
143 #a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR
144 [ * #ta * whd in ⊢ (%→?); >Hintape * * #c * whd in ⊢ (??%?→?); #Hc destruct (Hc) #Hx #Hta
145   * #tb * whd in ⊢ (%→?); * #Htb #_ cases (Htb (l1@〈a0,false〉::〈x0,true〉::l2) x) -Htb #Htb #_ lapply (Htb … Hta) -Hta -Htb #Htb
146   * #tc * whd in ⊢ (%→?); * #_ #Htc lapply (Htc … Htb) -Htb -Htc #Htc
147   * #td * whd in ⊢ (%→?); * #_ #Htd cases (Htd … Htc) -Htd #_ #Htd cases (Htd (refl …)) -Htd #Htd #_
148   lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ? (refl …) ?)
149   [#x1 #Hx1 cases (memb_append … Hx1) [@Hl1marks| -Hx1 #Hx1 >(memb_single … Hx1) % ]
150   |>associative_append % ] -Htd >reverse_append in ⊢ (???%→?); >associative_append in ⊢ (???%→?); #Htd
151   * #te * whd in ⊢ (%→?); * #Hte cases (Hte l2 x0) -Hte #Hte #_ #_ lapply (Hte … Htd) -Hte -Htd -Htc #Hte
152   * #tf * whd in ⊢ (%→?); * #_ #Htf lapply (Htf … Hte) -Hte -Htf
153   generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1)
154   [ #Hl1marks #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
155     [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
156     * * #_ #Houtc lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc
157     #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) #_ % % [%] @Houtc
158   | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single
159     #Hth whd in ⊢ (%→?); * #_ #Houtc cases (Houtc … Hth) -Houtc
160     [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ]
161     * * #Hc1 #Houtc #_ lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc
162     [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1
163       [@Hl1marks @memb_append_l1 <(reverse_reverse … l1) @memb_reverse @Hx1 
164       |>(memb_single … Hx1) % ]
165     | normalize >associative_append % ] 
166     #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc
167     >reverse_append >reverse_reverse >associative_append >associative_append % ]
168 | * #ta * whd in ⊢ (%→?); >Hintape * #Hxc
169   cut (x ≠ null) [ % #Hx cases (\Pf (Hxc ? (refl …))) #Hfalse @Hfalse >Hx % ] -Hxc #Hxnull
170   #Hta whd in ⊢ (%→?); #Houtc %2 % // <Hta @Houtc ]
171 qed.
172
173 definition copy_step ≝
174   ifTM ? (test_char STape (λc.bit_or_null (\fst c)))
175   (single_finalTM ? (copy_step_subcase FSUnialpha (bit false)
176     (copy_step_subcase FSUnialpha (bit true) nocopy_subcase)))
177   (nop ?)
178   tc_true.
179   
180 definition R_copy_step_true ≝ 
181   λt1,t2.
182     ∀ls,c,rs. t1 = midtape STape ls 〈c,true〉 rs → 
183     bit_or_null c = true ∧
184     (∀a,l1,x0,a0,l2,l3.
185      ls = (l1@〈a0,false〉::〈x0,true〉::l2) → 
186      rs = (〈a,false〉::l3) → 
187      no_marks l1 →          
188      ((∃x. c = bit x ∧
189       t2 = midtape STape (〈bit x,false〉::l1@〈a0,true〉::〈bit x,false〉::l2) 〈a,true〉 l3) ∨
190       (c = null ∧
191       t2 = midtape ? (〈null,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3))).
192      
193 definition R_copy_step_false ≝ 
194   λt1,t2.
195    ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → 
196    bit_or_null (\fst c) = false ∧ t2 = t1.
197    
198 lemma sem_copy_step : 
199   accRealize ? copy_step (inr … (inl … (inr … start_nop))) R_copy_step_true R_copy_step_false.
200 #intape
201 @(acc_sem_if_app … (sem_test_char ? (λc:STape.bit_or_null (\fst c)))  …
202     (sem_copy_step_subcase FSUnialpha (bit false) …
203        (sem_copy_step_subcase FSUnialpha (bit true) … (sem_nocopy_subcase …)))
204           (sem_nop …))
205 [ #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1 >Ht1 in H1;
206   * * #c0 * whd in ⊢ (??%?→?); #Hc0 destruct (Hc0) #Hc #Ht3 % //
207   #a #l1 #x0 #a0 #l2 #l3 #Hls #Hrs #Hl1marks >Hls in Ht3; >Hrs #Ht3
208   cases (H2 … Ht3 ?)
209   [ * #Hc' #Ht2 % %{false} % // <Hc' @Ht2
210   | * #Hnotfalse whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
211     [ * #Hc' #Ht2 % %{true} % // <Hc' @Ht2
212     |  * #Hnottrue whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
213       [ * #Hc' #Ht2 %2 <Hc' % // @Ht2
214       | * #Hnotnull @False_ind
215         generalize in match Hnotnull;generalize in match Hnottrue;generalize in match Hnotfalse;
216         cases c in Hc; normalize
217         [ * [ #_ #_ * #Hfalse #_ | #_ * #Hfalse #_ #_ ]
218         | #_ #_ #_ * #Hfalse
219         |*: #Hfalse destruct (Hfalse) ] @Hfalse %
220       | @Hl1marks ]
221     | @Hl1marks ]
222   | @Hl1marks ]
223 | #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1
224   >Ht1 in H1; * #Hcur #Ht3 % // @Hcur % ]
225 qed.
226
227 (*
228 1) il primo carattere è marcato
229 2) l'ultimo carattere è l'unico che può essere null, gli altri sono bit
230 3) il terminatore non è né bit, né null
231 *)
232    
233 definition copy0 ≝ whileTM ? copy_step (inr … (inl … (inr … start_nop))).
234
235 let rec merge_config (l1,l2:list STape) ≝ 
236   match l1 with
237   [ nil ⇒ nil ?
238   | cons p1 l1' ⇒ match l2 with
239     [ nil ⇒ nil ? 
240     | cons p2 l2' ⇒ 
241            let 〈c1,b1〉 ≝ p1 in let 〈c2,b2〉 ≝ p2 in
242            match c2 with
243            [ null ⇒ p1
244            | _ ⇒ p2 ] :: merge_config l1' l2' ] ].
245            
246 lemma merge_config_append :
247  ∀l1,l2,l3,l4.|l1| = |l2| → 
248  merge_config (l1@l3) (l2@l4) = merge_config l1 l2@merge_config l3 l4.
249 #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen)
250 [normalize //
251 | #t1 #t2 * #c1 #b1 * #c2 #b2 #IH whd in ⊢ (??%%); >IH % ]
252 qed.
253
254 definition R_copy0 ≝ λt1,t2.
255   ∀ls,c,c0,rs,l1,l3,l4.
256   t1 = midtape STape (l3@l4@〈c0,true〉::ls) 〈c,true〉 (l1@rs) → 
257   no_marks l1 → no_marks (l3@l4) → |l1| = |l4| → 
258   ∀l1',bv.〈c,false〉::l1 = l1'@[〈comma,bv〉] → only_bits_or_nulls l1' → 
259   ∀l4',bg.l4@[〈c0,false〉] = 〈grid,bg〉::l4' → only_bits_or_nulls l4' → 
260   (c = comma ∧ t2 = t1) ∨
261   (c ≠ comma ∧ 
262     t2 = midtape ? (reverse ? l1'@l3@〈grid,true〉::
263                   merge_config l4' (reverse ? l1')@ls) 
264      〈comma,true〉 rs).
265      
266 lemma inj_append_singleton_l1 :
267   ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → l1 = l2.
268 #A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
269 >reverse_append >reverse_append normalize #H1 destruct
270 lapply (eq_f … (reverse ?) … e0) >reverse_reverse >reverse_reverse //
271 qed.
272
273 lemma inj_append_singleton_l2 :
274   ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → a1 = a2.
275 #A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
276 >reverse_append >reverse_append normalize #H1 destruct %
277 qed.
278
279 axiom daemon : ∀P:Prop.P.
280
281 lemma list_cases2_full : 
282   ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop.
283   length ? l1 = length ? l2 →
284   (l1 = [] → l2 = [] → P) → 
285   (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P.
286 #T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl)
287 [ #Pnil #Pcons @Pnil //
288 | #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ]
289 qed.
290
291 lemma wsem_copy0 : WRealize ? copy0 R_copy0.
292 #intape #k #outc #Hloop 
293 lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop
294 * #ta * #Hstar @(star_ind_l ??????? Hstar)
295 [ whd in ⊢ (%→?); #Hleft
296   #ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
297   #Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits
298   cases (Hleft … Htb) -Hleft #Hc #Houtc % %
299   [ generalize in match Hl1bits; -Hl1bits cases l1' in Hl1;
300     [ normalize #Hl1 #c1 destruct (Hl1) %
301     | * #c' #b' #l0 #Heq normalize in Heq; destruct (Heq)
302       #Hl1bits lapply (Hl1bits 〈c',false〉 ?) [ @memb_hd ] 
303       >Hc #Hfalse destruct ]
304   | @Houtc ]
305 | #tc #td whd in ⊢ (%→?→(?→%)→%→?); #Htc #Hstar1 #Hind #Htd
306   lapply (Hind Htd) -Hind #Hind
307   #ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
308   #Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits %2
309   cases (Htc … Htb) -Htc #Hcbitnull #Htc
310   % [ % #Hc' >Hc' in Hcbitnull; normalize #Hfalse destruct (Hfalse) ]
311   cut (|l1| = |reverse ? l4|) [>length_reverse @Hlen] #Hlen1
312   @(list_cases2_full … Hlen1)
313   [ (* case l1 = [] is discriminated because l1 contains at least comma *)
314     #Hl1nil @False_ind >Hl1nil in Hl1; cases l1' normalize
315     [ #Hl1 destruct normalize in Hcbitnull; destruct (Hcbitnull)
316     | #p0 #l0 normalize #Hfalse destruct (Hfalse) cases l0 in e0;
317       [ normalize #Hfalse1 destruct (Hfalse1)
318       | #p0' #l0' normalize #Hfalse1 destruct (Hfalse1) ] ]
319   | (* case c::l1 = c::a::l1'' *)
320     * #a #ba * #a0 #ba0 #l1'' #l4'' #Hl1cons #Hl4cons
321     lapply (eq_f ?? (reverse ?) ?? Hl4cons) >reverse_reverse >reverse_cons -Hl4cons #Hl4cons
322     cut (ba = false) 
323     [ >Hl1cons in Hl1nomarks; #Hl1nomarks lapply (Hl1nomarks 〈a,ba〉 ?)
324       [ @memb_hd | normalize // ] ] #Hba
325     cut (ba0 = false) 
326     [ >Hl4cons in Hl3l4nomarks; #Hl3l4nomarks lapply (Hl3l4nomarks 〈a0,ba0〉 ?)
327       [ @memb_append_l2 @memb_append_l2 @memb_hd | normalize // ] ] #Hba0
328     >Hba0 in Hl4cons; >Hba in Hl1cons; -Hba0 -Hba #Hl1cons #Hl4cons
329     >Hl4cons in Htc; >Hl1cons #Htc
330     lapply (Htc a (l3@reverse ? l4'') c0 a0 ls (l1''@rs) ? (refl ??) ?)
331     [ #x #Hx @Hl3l4nomarks >Hl4cons <associative_append
332       @memb_append_l1 @Hx
333     | >associative_append >associative_append %
334     | -Htc
335       cut (∃la.l1' = 〈c,false〉::la)
336       [ >Hl1cons in Hl1; cases l1'
337         [normalize #Hfalse destruct (Hfalse)
338         | #p #la normalize #Hla destruct (Hla) @(ex_intro ?? la) % ] ]
339       * #la #Hla
340       cut (∃lb.l4' = lb@[〈c0,false〉])
341       [ >Hl4cons in Hl4;
342         @(list_elim_left … l4')
343         [ #Heq lapply (eq_f … (length ?) … Heq)
344           >length_append >length_append 
345           >commutative_plus normalize >commutative_plus normalize
346           #Hfalse destruct
347         | #a1 #tl #_ #Heq 
348           >(inj_append_singleton_l2 ? (reverse ? l4''@[〈a0,false〉]) (〈grid,bg〉::tl) 〈c0,false〉 a1 Heq)
349           @ex_intro //
350       ] ] * #lb #Hlb
351       cut (|lb| = |reverse ? la|) 
352       [ >Hla in Hl1; >Hlb in Hl4; #Hl4 #Hl1
353         >(?:l1 = la@[〈comma,bv〉]) in Hlen;
354         [|normalize in Hl1; destruct (Hl1) %]
355         >(?:l4 = 〈grid,bg〉::lb)
356         [|@(inj_append_singleton_l1 ?? (〈grid,bg〉::lb) ?? Hl4) ]
357         >length_append >commutative_plus >length_reverse
358         normalize #Hlalb destruct (Hlalb) //
359       ] #Hlen2
360       *
361       (* by hyp on the first iteration step, 
362          we consider whether c = bit x or c = null *)
363       (* c = bit x *)
364       [ * #x * #Hx #Htc 
365         lapply (Hind (〈bit x,false〉::ls) a a0 rs l1'' 
366                 (〈bit x,false〉::l3) (reverse ? l4'') ????) 
367         [ >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus 
368           normalize #Hlen destruct (Hlen) //
369         | #x0 #Hx0 cases (orb_true_l … Hx0)
370           [ #Hx0eq >(\P Hx0eq) %
371           | -Hx0 #Hx0 @Hl3l4nomarks >Hl4cons
372             <associative_append @memb_append_l1 // ]
373         | #x0 #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
374         | >Htc >associative_append % 
375         | -Hind 
376           <Hl1cons <Hl4cons #Hind lapply (Hind la bv ?? lb bg ??)
377           [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
378           | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
379             @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
380           | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
381           | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
382             destruct (Hl1) // ] -Hind
383           (* by IH, we proceed by cases, whether a = comma 
384              (consequently several lists = []) or not *)          
385           *
386           [ * #Ha #Houtc1
387            cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = [])
388            [ @daemon ] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
389            >Hl1cons in Hl1; >Hla
390            >Houtc1 >Htc #Hl1
391            >Hl4cons in Hl4; >Hlb #Hl4
392            >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hx
393            cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
394            normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
395           | * #Ha #Houtc1 >Houtc1 @eq_f3 //
396             >Hla >reverse_cons >associative_append @eq_f
397             >Hx whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
398             >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %            
399           ]
400        ]
401     | (* c = null *)
402       * #Hc #Htc 
403       lapply (Hind (〈c0,false〉::ls) a a0 rs l1'' (〈null,false〉::l3) (reverse ? l4'') ????)
404       [  >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus normalize
405          #Hlen destruct (Hlen) @e0
406       | #x0 #Hx0 cases (memb_append STape ? [〈null,false〉] (l3@reverse ? l4'') … Hx0) -Hx0 #Hx0
407         [ >(memb_single … Hx0) %
408         | @Hl3l4nomarks cases (memb_append … Hx0) -Hx0 #Hx0
409           [ @memb_append_l1 //
410           | @memb_append_l2 >Hl4cons @memb_append_l1 // ]
411         ]
412       | >Hl1cons #x' #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
413       | >Htc @eq_f3 // >associative_append % ] -Hind <Hl1cons <Hl4cons #Hind
414         lapply (Hind la bv ?? lb bg ??)
415           [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
416           | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
417             @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
418           | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
419           | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
420             destruct (Hl1) // ] -Hind *
421           (* by IH, we proceed by cases, whether a = comma 
422              (consequently several lists = []) or not *)          
423           [ * #Ha #Houtc1 >Hl1cons in Hl1; >Hla
424            >Houtc1 >Htc #Hl1
425            >Hl4cons in Hl4; >Hlb #Hl4
426            cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = []) 
427            [@daemon] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
428            >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hc
429            cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
430            normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
431           | * #Ha #Houtc1 >Houtc1 @eq_f3 //
432             >Hla >reverse_cons >associative_append @eq_f
433             >Hc whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
434             >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %
435           ]
436        ]
437 ]]]
438 qed.
439
440 definition Pre_copy0 ≝ λt1.
441   ∃ls,c,c0,rs,l1,l3,l4.
442   t1 = midtape STape (l3@l4@〈c0,true〉::ls) 〈c,true〉 (l1@rs) ∧
443   no_marks l1 ∧ no_marks (l3@l4) ∧ |l1| = |l4| ∧
444   (∃l1',bv.〈c,false〉::l1 = l1'@[〈comma,bv〉] ∧ only_bits_or_nulls l1') ∧
445   (∃l4',bg.l4@[〈c0,false〉] = 〈grid,bg〉::l4' ∧ only_bits_or_nulls l4').
446
447 definition Pre_copy ≝ λt1.
448   ∃ls,s0,s1,c0,c1,rs,l1,l3,l4.
449   t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) ∧
450   no_marks l1 ∧ no_marks l3 ∧ no_marks l4 ∧ |l1| = |l4| ∧ 
451   only_bits (l4@[〈s0,false〉]) ∧ only_bits (〈s1,false〉::l1) ∧ 
452   bit_or_null c0 = true ∧ bit_or_null c1 = true.
453   
454 lemma list_last: ∀A.∀l:list A.
455   l = [ ] ∨ ∃a,l1. l = l1@[a].
456 #A #l <(reverse_reverse ? l) cases (reverse A l)
457   [%1 //
458   |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
459   ]
460 qed.
461
462 lemma terminate_copy0 : ∀t.Pre_copy0 t → Terminate ? copy0 t.
463 #t #HPre
464 @(terminate_while_guarded ??? 
465    Pre_copy0 … 
466    (acc_Realize_to_acc_GRealize ??? Pre_copy0 … sem_copy_step)
467    … HPre) [%]
468   [ -HPre -t #t1 #t2 * #ls * #c * #c0 * #rs * #l1 * #l3 * #l4
469     * * * * * #Ht1 #Hl1nomarks #Hl3l4nomarks #Hlen
470     * #l1' * #bv * #Hl1 #Hbitsnullsl1' * #l4' * #bg * #Hl4 #Hbitsnullsl4'
471     #HR cases (HR … Ht1) -HR #Hbitnullc
472     cut (∃d1,l1''.l1 = 〈d1,false〉::l1'')
473     [ lapply Hl1nomarks cases l1 in Hl1;
474       [ whd in ⊢ (???%→?); #Hl1
475         lapply (append_l2_injective_r ? [] l1' [〈c,false〉] [〈comma,bv〉] (refl ??) Hl1)
476         #Hc destruct (Hc) normalize in Hbitnullc; destruct (Hbitnullc) 
477       | * #d #bd #l1'' #_ #Hl1nomarks >(?:bd = false) [| @(Hl1nomarks 〈d,bd〉) @memb_hd ] /3/ ] ]
478     * #d1 * #l1'' #Hl1''
479     cut (∃d2,l4''.l4 = l4''@[〈d2,false〉])
480     [ lapply Hl4 cases (list_last ? l4)
481       [ #Hl4' >Hl4' in Hlen; >Hl1'' normalize in ⊢ (%→?); #HFalse destruct (HFalse)
482       | * * #d2 #bd2 * #l4'' #Hl4'' >Hl4'' in Hl3l4nomarks; #Hl3l4nomarks #_ 
483         <(?:bd2 = false) [| @(Hl3l4nomarks 〈d2,bd2〉) @memb_append_l2 @memb_append_l2 @memb_hd ]
484         /3/ ] ]
485     * #d2 * #l4'' #Hl4'' >Hl1'' >Hl4''
486     cut (∃l1''',bv0.〈d1,false〉::l1''=l1'''@[〈comma,bv0〉]∧only_bits_or_nulls l1''')
487     [ <Hl1'' cases (list_last ? l1) in Hl1'';
488         [ #Hl1'' >Hl1''#HFalse destruct(HFalse)
489         | * #a * #l1''' #Hl1''' >Hl1''' in Hl1; #Hl1 #_
490           lapply (append_l2_injective_r ? (〈c,false〉::l1''') l1' [a] [〈comma,bv〉] (refl …) Hl1)
491           #Ha destruct (Ha) %{l1'''} %{bv} % //
492           #x #Hx @Hbitsnullsl1'
493           lapply (append_l1_injective_r ? (〈c,false〉::l1''') l1' [〈comma,bv〉] [〈comma,bv〉] (refl …) Hl1)
494           #H <H @memb_cons @Hx ] ]
495     cut (∃l4''',bg0.l4''@[〈d2,false〉]=〈grid,bg0〉::l4'''∧only_bits_or_nulls l4''')
496     [ <Hl4'' cases (list_last ? l4') in Hl4;
497         [ #Hl4' >Hl4' >Hl4'' cases l4''
498           [ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
499           | #y #yl normalize in ⊢ (%→?); #H1 destruct (H1) cases yl in e0;
500             [ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
501             | #z #zl normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] ]
502         | * #a * #l4''' #Hl4''' >Hl4''' #Hl4
503           lapply (append_l1_injective_r ? l4 (〈grid,bg〉::l4''') [〈c0,false〉] [a] (refl …) Hl4)
504           -Hl4 #Hl4 >Hl4 %{l4'''} %{bg} % //
505           #x #Hx @Hbitsnullsl4' >Hl4''' @memb_append_l1 @Hx ] ]
506     #Hl4''' #Hl1'''
507     #HR cases (HR d1 (l3@l4'') c0 d2 ls (l1''@rs) ? ? ?)
508     [3: >associative_append >associative_append %
509     |4: %
510     |5: #x #Hx @Hl3l4nomarks cases (memb_append … Hx)
511       [ @memb_append_l1
512       | >Hl4'' -Hx #Hx @memb_append_l2 @memb_append_l1 @Hx ]
513     | * #x1 * #Hc #Ht2 whd >Ht2
514       %{(〈bit x1,false〉::ls)} %{d1} %{d2} %{rs} %{l1''} %{(〈bit x1,false〉::l3)} %{l4''}
515       % [ % [ % [ % [ % 
516       [ >associative_append %
517       | #x #Hx @Hl1nomarks >Hl1'' @memb_cons @Hx ]
518       | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
519         [ >(\P Hx) %
520         | @Hl3l4nomarks cases (memb_append … Hx)
521           [ @memb_append_l1
522           | -Hx #Hx >Hl4'' @memb_append_l2 @memb_append_l1 @Hx ] ] ]
523       | >Hl1'' in Hlen; >Hl4'' >length_append >commutative_plus
524         normalize in ⊢ (%→?); #Hlen destruct (Hlen) @e0 ]
525       | @Hl1''' ]
526       | @Hl4''' ]
527    | * #Hc #Ht2 whd >Ht2
528       %{(〈c0,false〉::ls)} %{d1} %{d2} %{rs} %{l1''} %{(〈null,false〉::l3)} %{l4''}
529       % [ % [ % [ % [ % 
530       [ >associative_append %
531       | #x #Hx @Hl1nomarks >Hl1'' @memb_cons @Hx ]
532       | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
533         [ >(\P Hx) %
534         | @Hl3l4nomarks cases (memb_append … Hx)
535           [ @memb_append_l1
536           | -Hx #Hx >Hl4'' @memb_append_l2 @memb_append_l1 @Hx ] ] ]
537       | >Hl1'' in Hlen; >Hl4'' >length_append >commutative_plus
538         normalize in ⊢ (%→?); #Hlen destruct (Hlen) @e0 ]
539       | @Hl1''' ] 
540       | @Hl4''' ]
541     ]
542  |cases HPre -HPre #ls * #c * #c0 * #rs * #l1 * #l3 * #l4
543   * * * * * #Ht #Hl1nomarks #Hl3l4nomarks #Hlen
544   * #l1' * #bv * #Hl1 #Hbitsnullsl1'
545   * #l4' * #bg * #Hl4 #Hbitsnullsl4'
546   >Ht lapply Hbitsnullsl1' lapply Hl1 lapply l1' lapply Hl3l4nomarks 
547   lapply Hl1nomarks lapply l3 lapply c0 lapply c lapply ls 
548   -Hbitsnullsl1' -Hl1 -l1' -Hl3l4nomarks -Hl1nomarks -l3 -Hl4 -c0 -c -ls
549   <(reverse_reverse ? l4) <(length_reverse ? l4) in Hlen; #Hlen
550   @(list_ind2 … Hlen)
551   [ #ls #c #c0 #l3 #_ #_ #l1' #Hl1 #Hbitsnullsl1' cases l1' in Hl1;
552     [| #x * [| #x0 #xs ] normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
553     normalize in ⊢ (%→?); #Hl1' destruct (Hl1') % #t1 whd in ⊢ (%→?);
554     #HR cases (HR … (refl …)) whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
555   | #xs #ys * #x #bx * #y #by #IH
556     #ls #c #c0 #l3 #Hl1nomarks #Hl3l4nomarks #l1' #Hl1 
557     #Hbitsnullsl1' %
558     #t1 whd in ⊢ (%→?); #HR cases (HR … (refl …)) -HR
559     #Hbitnullc #HR
560     lapply (Hl1nomarks 〈x,bx〉 (memb_hd …)) normalize in ⊢ (%→?); #Hbx destruct (Hbx)
561     (*cut (∃d2,l4''.〈y,by〉::ys = l4''@[〈d2,false〉])
562     [ cases (list_last ? ys) in Hl3l4nomarks;
563       [ #Hl4' #Hl3l4nomarks >Hl4' >(?:by = false) [%{y} %{([])} %]
564         @(Hl3l4nomarks 〈y,by〉) @memb_append_l2 @memb_hd
565       | * * #d2 #bd2 * #l4'' #Hl4'' >Hl4'' #Hl3l4nomarks 
566         <(?:bd2 = false) [| @(Hl3l4nomarks 〈d2,bd2〉) @memb_append_l2 @memb_cons @memb_append_l2 @memb_hd ]
567         %{d2} %{(〈y,by〉::l4'')} % ] ]
568     * #a0 * #l4'' #Hl4'' >Hl4'' in HR; #HR *)
569     cases (HR x (l3@reverse ? ys) c0 y ls (xs@rs) ? (refl …) ?)
570     [3: >reverse_cons >associative_append >associative_append
571         >(?:by = false) [|@(Hl3l4nomarks 〈y,by〉) @memb_append_l2 >reverse_cons @memb_append_l2 @memb_hd ] %
572     |4: #x #Hx @Hl3l4nomarks >reverse_cons <associative_append @memb_append_l1 @Hx
573     | * #x1 * #Hc #Ht1 >Ht1 
574       <(?: (〈bit x1,false〉::l3)@reverse (FSUnialpha×bool) ys@〈y,true〉::〈bit x1,false〉::ls =
575             〈bit x1,false〉::(l3@reverse (FSUnialpha×bool) ys)@〈y,true〉::〈bit x1,false〉::ls)
576      [cut (∃l1''.〈x,false〉::xs = l1''@[〈comma,bv〉] ∧ only_bits_or_nulls l1'')
577       [lapply Hbitsnullsl1' lapply Hl1 -Hl1 -Hbitsnullsl1' cases l1'
578        [normalize in ⊢(%→?); #Hfalse destruct (Hfalse)
579        |#x2 #l1'' normalize in ⊢ (%→?); #Heq #Hbitsnullsl1' destruct (Heq)
580         %{l1''} % [@e0]
581         #x #Hx @Hbitsnullsl1' @memb_cons @Hx ] ]
582       * #l1'' * #Hl1'' #Hbitsnullsl1''
583       @(IH (〈bit x1,false〉::ls) x y (〈bit x1,false〉::l3) ?? l1'' Hl1'' Hbitsnullsl1'')
584       [#x #Hx @Hl1nomarks @memb_cons @Hx
585       |#x #Hx cases (orb_true_l … Hx) -Hx #Hx
586        [>(\P Hx) %
587        |cases (memb_append … Hx) -Hx #Hx @Hl3l4nomarks
588         [@memb_append_l1 @Hx
589         |@memb_append_l2 >reverse_cons @memb_append_l1 @Hx ] ] ]
590      |>associative_append % ]
591     | * #Hc #Ht1 >Ht1 
592       <(?: (〈null,false〉::l3)@reverse (FSUnialpha×bool) ys@〈y,true〉::〈c0,false〉::ls =
593             〈null,false〉::(l3@reverse (FSUnialpha×bool) ys)@〈y,true〉::〈c0,false〉::ls)
594      [cut (∃l1''.〈x,false〉::xs = l1''@[〈comma,bv〉] ∧ only_bits_or_nulls l1'')
595       [lapply Hbitsnullsl1' lapply Hl1 -Hl1 -Hbitsnullsl1' cases l1'
596        [normalize in ⊢(%→?); #Hfalse destruct (Hfalse)
597        |#x2 #l1'' normalize in ⊢ (%→?); #Heq #Hbitsnullsl1' destruct (Heq)
598         %{l1''} % [@e0]
599         #x #Hx @Hbitsnullsl1' @memb_cons @Hx ] ]
600       * #l1'' * #Hl1'' #Hbitsnullsl1''
601       @(IH (〈c0,false〉::ls) x y (〈null,false〉::l3) ?? l1'' Hl1'' Hbitsnullsl1'')
602       [#x #Hx @Hl1nomarks @memb_cons @Hx
603       |#x #Hx cases (orb_true_l … Hx) -Hx #Hx
604        [>(\P Hx) %
605        |cases (memb_append … Hx) -Hx #Hx @Hl3l4nomarks
606         [@memb_append_l1 @Hx
607         |@memb_append_l2 >reverse_cons @memb_append_l1 @Hx ] ] ]
608     |>associative_append % ]
609 ]]]
610 qed.
611
612 definition merge_char ≝ λc1,c2.
613   match c2 with
614   [ null ⇒ c1
615   | _ ⇒ c2 ].
616   
617 lemma merge_cons : 
618   ∀c1,c2,conf1,conf2.
619   merge_config (〈c1,false〉::conf1) (〈c2,false〉::conf2) = 
620     〈merge_char c1 c2,false〉::merge_config conf1 conf2.
621 #c1 #c2 #conf1 #conf2 normalize @eq_f2 //
622 cases c2 /2/
623 qed.
624
625 lemma merge_bits : ∀l1,l2.|l1| = |l2| → only_bits l2 → merge_config l1 l2 = l2.
626 #l1 #l2 #Hlen @(list_ind2 … Hlen) //
627 #tl1 #tl2 #hd1 #hd2 #IH
628 >(eq_pair_fst_snd … hd1) >(eq_pair_fst_snd … hd2) #Hbits
629 change with (cons ???) in ⊢ (??%?); @eq_f2
630 [ cases (\fst hd2) in Hbits; 
631   [ #b #_ %
632   |*: #Hfalse lapply (Hfalse … (memb_hd …)) normalize #Hfalse1 destruct (Hfalse1) ]
633 | @IH #x #Hx @Hbits @memb_cons // ]
634 qed.
635
636 lemma merge_config_c_nil : 
637   ∀c.merge_config c [] = [].
638 #c cases c normalize //
639 qed.
640
641 lemma reverse_merge_config :
642   ∀c1,c2.|c1| = |c2| → reverse ? (merge_config c1 c2) = 
643     merge_config (reverse ? c1) (reverse ? c2).        
644 #c1 #c2 <(length_reverse ? c1) <(length_reverse ? c2) #Hlen
645 <(reverse_reverse ? c1) in ⊢ (??%?); <(reverse_reverse ? c2) in ⊢ (??%?);
646 generalize in match Hlen; @(list_ind2 … Hlen) -Hlen //
647 #tl1 #tl2 #hd1 #hd2 #IH whd in ⊢ (??%%→?); #Hlen destruct (Hlen) -Hlen
648 <(length_reverse ? tl1) in e0; <(length_reverse ? tl2) #Hlen
649 >reverse_cons >reverse_cons >(merge_config_append ???? Hlen)
650 >reverse_append >(eq_pair_fst_snd ?? hd1) >(eq_pair_fst_snd ?? hd2)
651 whd in ⊢ (??%%); @eq_f2 // @IH //
652 qed.
653
654 definition copy
655 ≝ 
656   seq STape copy0 (seq ? (move_l …) (seq ? (adv_to_mark_l … (is_marked ?))
657    (seq ? (clear_mark …) (seq ? (adv_to_mark_r … (is_marked ?)) (clear_mark …))))).
658
659 (*
660    s0, s1 = caratteri di testa dello stato
661    c0 = carattere corrente del nastro oggetto
662    c1 = carattere in scrittura sul nastro oggetto
663    
664    questa dimostrazione sfrutta il fatto che 
665    merge_config (l0@[c0]) (l1@[c1]) = l1@[merge_char c0 c1] 
666    se l0 e l1 non contengono null
667 *)
668
669 definition R_copy ≝ λt1,t2.
670   ∀ls,s0,s1,c0,c1,rs,l1,l3,l4.
671   t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) → 
672   no_marks l1 → no_marks l3 → no_marks l4 → |l1| = |l4| → 
673   only_bits (l4@[〈s0,false〉]) → only_bits (〈s1,false〉::l1) → 
674   bit_or_null c0 = true → bit_or_null c1 = true → 
675   t2 = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l3@〈grid,false〉::
676                       〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
677        〈comma,false〉 rs.
678        
679 lemma sem_copy0 : GRealize ? copy0 Pre_copy0 R_copy0.
680 @WRealize_to_GRealize
681 [ @terminate_copy0
682 | @wsem_copy0 ]
683 qed.
684
685 definition option_cons ≝ λA.λa:option A.λl.
686   match a with
687   [ None ⇒ l
688   | Some a' ⇒ a'::l ].
689   
690 lemma sem_copy : GRealize ? copy Pre_copy R_copy.
691 @(GRealize_to_GRealize_2 ?? Pre_copy0 ? R_copy) //
692 (* preamble: Pre_copy0 implies Pre_copy *)
693 [ #t1 * #ls * #s0 * #s1 * #c0 * #c1 * #rs * #l1 * #l3 * #l4 * * * * * * * *
694   #Ht1 #Hl1nomarks #Hl3nomarks #Hl4nomarks #Hlen #Hbitsl4 #Hbitsl1 
695   #Hbitnullc0 #Hbitnullc1 whd
696   %{ls} %{s1} %{s0} %{rs} %{(l1@[〈c1,false〉;〈comma,false〉])}
697   %{l3} %{(〈grid,false〉::〈c0,false〉::l4)}
698   % [ % [ % [ % [ % 
699   [ >Ht1 -Ht1 @eq_f >associative_append %
700   | #x #Hx cases (memb_append … Hx) -Hx #Hx
701     [ @(Hl1nomarks ? Hx)
702     | cases (orb_true_l … Hx) -Hx #Hx 
703       [ >(\P Hx) %
704       | >(memb_single … Hx) % ] ] ]
705   | #x #Hx cases (memb_append … Hx) -Hx #Hx
706     [ @(Hl3nomarks ? Hx)
707     | cases (orb_true_l … Hx) -Hx #Hx
708       [ >(\P Hx) %
709       | cases (orb_true_l … Hx) -Hx #Hx
710         [ >(\P Hx) %
711         | @(Hl4nomarks ? Hx) ]]]]
712   |>length_append >Hlen >commutative_plus % ]
713   | %{(〈s1,false〉::l1@[〈c1,false〉])} %{false} %
714     [ @eq_f >associative_append %
715     | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
716       [ >(\P Hx) @is_bit_to_bit_or_null @(Hbitsl1 〈s1,false〉) @memb_hd
717       | cases (memb_append … Hx) -Hx #Hx
718         [ @is_bit_to_bit_or_null @(Hbitsl1 〈\fst x,\snd x〉) @memb_cons <eq_pair_fst_snd @Hx
719         | >(memb_single… Hx) // ]
720       ]
721     ] ]
722   | %{(〈c0,false〉::l4@[〈s0,false〉])} %{false} %
723     [ %
724     | #x #Hx cases (orb_true_l … Hx) -Hx #Hx
725       [ >(\P Hx) //
726       | cases (memb_append … Hx) -Hx #Hx
727         [ @is_bit_to_bit_or_null @(Hbitsl4 〈\fst x,\snd x〉) @memb_append_l1 <eq_pair_fst_snd @Hx
728         | >(memb_single… Hx) @is_bit_to_bit_or_null @(Hbitsl4 〈s0,false〉) @memb_append_l2 @memb_hd ]
729       ]
730 ] ] ]
731 (*whd in ⊢ (%→%)
732 change with (?·?) in match copy;*)
733 @(sem_seq_app_guarded ???? (λx.True) ??? sem_copy0)
734 [2:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_move_l (FinProd FSUnialpha FinBool))))
735   [@(λx.True)
736   |4://
737   |5:@sub_reflexive
738   |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_adv_to_mark_l … (is_marked FSUnialpha))))
739     [@(λx.True)
740     |4://
741     |5:@sub_reflexive
742     |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_clear_mark ?)))
743       [@(λx.True)
744       |4://
745       |5:@sub_reflexive
746       |3:@(sem_seq_app_guarded ???????? (Realize_to_GRealize … (sem_adv_to_mark_r ? (is_marked ?))))
747         [@(λx.True)
748         |4://
749         |5:@sub_reflexive
750         |3:@(Realize_to_GRealize … (sem_clear_mark ?)) ] ] ] ]
751 |3: //
752 |1:skip]
753 #intape #outtape #HR
754 #ls #s0 #s1 #c0 #c1 #rs #l1 #l2 #l3 #Hintape #Hl1marks #Hl2marks #Hl3marks #Hlen
755 #Hbits1 #Hbits2 #Hc0bits #Hc1bits
756 cases HR -HR #ta * whd in ⊢ (%→?); #Hta 
757 cut (ta = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l2@〈grid,true〉::
758                       〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
759        〈comma,true〉 rs)
760 [lapply (Hta ls s1 s0 rs (l1@[〈c1,false〉;〈comma,false〉]) l2 (〈grid,false〉::〈c0,false〉::l3) ?)
761   [>associative_append in ⊢ (???(????%)); normalize in ⊢ (???(??%%%)); @Hintape ]
762  -Hta #Hta cases (Hta ??? (〈s1,false〉::l1@[〈c1,false〉]) false ? ? ?? (refl ??) ?)
763   [3: #x #Hx cases (memb_append … Hx) -Hx #Hx
764     [ @Hl1marks //
765     | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | >(memb_single … Hx) % ]] 
766   |4: #x #Hx cases (memb_append … Hx) -Hx #Hx
767     [ @Hl2marks //
768     | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | cases (orb_true_l … Hx) [-Hx #Hx >(\P Hx) % | @Hl3marks ] ] ]
769   |5: >length_append normalize >Hlen >commutative_plus %
770   |6: normalize >associative_append %
771   |7: #x #Hx cases (memb_append ?? (〈s1,false〉::l1) … Hx) -Hx #Hx
772     [ whd in ⊢ (??%?); >(Hbits2 … Hx) %
773     | >(memb_single … Hx) // ]
774   |8: #x #Hx cases (memb_append … Hx) -Hx #Hx
775     [ cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) // | whd in ⊢ (??%?); >Hbits1 // @memb_append_l1 // ]
776     | >(memb_single … Hx) whd in ⊢ (??%?); >(Hbits1 〈s0,false〉) // @memb_append_l2 @memb_hd ]
777   | * #Hs1 @False_ind >Hs1 in Hbits2; #Hbits2 lapply (Hbits2 〈comma,false〉 ?) //
778     normalize #Hfalse destruct (Hfalse)
779   | * #Hs1 #Ht2 >Ht2 >reverse_cons >reverse_append >reverse_single @eq_f3 //
780     >merge_cons >merge_bits
781     [2: #x #Hx @Hbits2 cases (memb_append STape ? (reverse ? l1) ? Hx) -Hx #Hx
782       [@daemon | >(memb_single … Hx) @memb_hd ]
783     |3: >length_append >length_append >length_reverse >Hlen % ]
784     normalize >associative_append normalize >associative_append %
785   ]
786 ] -Hta #Hta * #tb * * #_ #Htb
787 lapply (Htb … Hta) -Htb #Htb change with (midtape ????) in Htb:(???%);
788 * #tc * * #_ #Htc 
789 cases (Htc … Htb)
790 (* [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] *)
791 #_ #Htc cases (Htc (refl ??)) -Htc #Htc #_
792 lapply (Htc (reverse ? l1@〈s1,false〉::l2) 〈grid,true〉
793           (〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)???)
794 [ #x #Hx cases (memb_append … Hx) -Hx #Hx
795   [ @Hl1marks @daemon
796   | cases (orb_true_l … Hx) -Hx #Hx
797     [ >(\P Hx) % | @(Hl2marks … Hx) ] ]
798 | %
799 | whd in ⊢ (??%?); >associative_append % ] -Htc #Htc
800 * #td * whd in ⊢ (%→?); * #_ #Htd lapply (Htd … Htc) -Htd #Htd
801 * #te * whd in ⊢ (%→?); * #_ #Hte cases (Hte … Htd) -Hte -Htd
802 [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
803 * * #_ #Hte #_ 
804 lapply (Hte (reverse ? (reverse ? l1@〈s1,false〉::l2)@[〈c1,false〉])
805          〈comma,true〉 rs ? (refl ??) ?) -Hte
806 [ >reverse_append >reverse_cons >reverse_reverse #x #Hx
807   cases (memb_append … Hx) -Hx #Hx
808   [ cases (memb_append … Hx) -Hx #Hx
809     [ cases (memb_append … Hx) -Hx #Hx
810       [ @daemon 
811       | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
812     | @(Hl1marks … Hx) ]
813   | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
814 | >reverse_append >reverse_reverse >reverse_cons
815   >associative_append >associative_append >associative_append
816   >associative_append >associative_append % ]
817 #Hte * #_ #Houtc lapply (Houtc … Hte) -Houtc #Houtc >Houtc
818 @eq_f3 //
819 >reverse_append >reverse_append >reverse_single >reverse_cons
820 >reverse_append >reverse_append >reverse_reverse >reverse_reverse
821 >reverse_single >associative_append >associative_append %
822 qed.