]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/universal/copy.ma
work in progress
[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 lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc
77     >reverse_append >reverse_reverse >associative_append >associative_append % ]
78 | * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta ? (refl ??)) -Hta 
79   #Hxc #Hta >Hta #Houtc %2 % // lapply (\Pf Hxc) @not_to_not #Heq >Heq % ]
80 qed.
81     
82 (*
83 if current = 0,tt
84    then advance_mark_r;
85         move_l;
86         advance_to_mark_l;
87         write(0,ff)
88         move_r;
89         mark;
90         move_r;
91         advance_to_mark_r;
92 else if current = 1,tt
93    then advance_mark_r;
94         move_l;
95         advance_to_mark_l;
96         write(1,ff)
97         move_r;
98         mark;
99         move_r;
100         advance_to_mark_r;
101 else if current = null 
102    then advance_mark_r;
103         move_l;
104         advance_to_mark_l
105         adv_mark_r;
106         move_r;
107         advance_to_mark_r
108 *)
109
110 definition nocopy_subcase ≝
111   ifTM STape (test_char ? (λx:STape.x == 〈null,true〉))
112     (seq ? (adv_mark_r …)
113       (seq ? (move_l …)
114         (seq ? (adv_to_mark_l … (is_marked ?))
115           (seq ? (adv_mark_r …)
116             (seq ? (move_r …) (adv_to_mark_r … (is_marked ?)))))))
117     (nop ?) tc_true.
118
119 definition R_nocopy_subcase ≝ 
120   λt1,t2.
121     ∀a,l1,x0,a0,l2,x,l3.
122     t1 = midtape STape (l1@〈a0,false〉::〈x0,true〉::l2) 
123          〈x,true〉 (〈a,false〉::l3) → 
124     (∀c.memb ? c l1 = true → is_marked ? c = false) →          
125     (x = null ∧
126      t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3) ∨
127     (x ≠ null ∧ t2 = t1).
128     
129 lemma sem_nocopy_subcase : Realize ? nocopy_subcase R_nocopy_subcase.
130 #intape
131 cases (sem_if ? (test_char ? (λx:STape.x == 〈null,true〉)) ?????? tc_true
132         (sem_test_char ? (λx:STape.x == 〈null,true〉))
133           (sem_seq … (sem_adv_mark_r …)
134             (sem_seq … (sem_move_l …)
135               (sem_seq … (sem_adv_to_mark_l … (is_marked ?))
136                 (sem_seq … (sem_adv_mark_r …)
137                   (sem_seq … (sem_move_r …) 
138                     (sem_adv_to_mark_r … (is_marked ?))))))) (sem_nop ?) intape)
139 #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc)  % [@Hloop] -Hloop
140 #a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR
141 [ * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta … (refl ??)) -Hta #Hx #Hta
142   * #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) -Hta -Htb #Htb
143   * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htb -Htc #Htc
144   * #td * whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd
145   [ >Htc * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
146   * #_ #Htd lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ? (refl ??) ?) -Htd
147   [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1 [@(Hl1marks ? Hx1)|>(memb_single … Hx1) %]
148   | normalize >associative_append % ] >reverse_append #Htd
149   * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd -Htc #Hte
150   * #tf * whd in ⊢ (%→?); #Htf lapply (Htf … Hte) -Hte -Htf
151   generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1)
152   [ #Hl1marks #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc
153     [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
154     * #_ #Houtc lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc
155     #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] @Houtc
156   | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single
157     #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc
158     [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ]
159     * #_ #Houtc lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc
160     [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1
161       [@Hl1marks @memb_append_l1 <(reverse_reverse … l1) @memb_reverse @Hx1 
162       |>(memb_single … Hx1) % ]
163     | normalize >associative_append % ] 
164     #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc
165     >reverse_append >reverse_reverse >associative_append >associative_append % ]
166 | * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta ? (refl ??)) -Hta 
167   #Hxc #Hta >Hta whd in ⊢ (%→?); #Houtc %2 %
168   [ lapply (\Pf Hxc) @not_to_not #Heq >Heq %
169   | @Houtc ]
170 qed.
171
172 definition copy_step ≝
173   ifTM ? (test_char STape (λc.bit_or_null (\fst c)))
174   (single_finalTM ? (copy_step_subcase FSUnialpha (bit false)
175     (copy_step_subcase FSUnialpha (bit true) nocopy_subcase)))
176   (nop ?)
177   tc_true.
178   
179 definition R_copy_step_true ≝ 
180   λt1,t2.
181     ∀ls,c,rs. t1 = midtape STape ls 〈c,true〉 rs → 
182     bit_or_null c = true ∧
183     (∀a,l1,x0,a0,l2,l3.
184      ls = (l1@〈a0,false〉::〈x0,true〉::l2) → 
185      rs = (〈a,false〉::l3) → 
186      no_marks l1 →          
187      ((∃x. c = bit x ∧
188       t2 = midtape STape (〈bit x,false〉::l1@〈a0,true〉::〈bit x,false〉::l2) 〈a,true〉 l3) ∨
189       (c = null ∧
190       t2 = midtape ? (〈null,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3))).
191      
192 definition R_copy_step_false ≝ 
193   λt1,t2.
194    ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → 
195    bit_or_null (\fst c) = false ∧ t2 = t1.
196    
197 lemma sem_copy_step : 
198   accRealize ? copy_step (inr … (inl … (inr … start_nop))) R_copy_step_true R_copy_step_false.
199 #intape
200 @(acc_sem_if_app … (sem_test_char ? (λc:STape.bit_or_null (\fst c)))  …
201     (sem_copy_step_subcase FSUnialpha (bit false) …
202        (sem_copy_step_subcase FSUnialpha (bit true) … (sem_nocopy_subcase …)))
203           (sem_nop …))
204 [ #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1 >Ht1 in H1; #H1
205   cases (H1 … (refl ??)) #Hc #Ht3 % [ @Hc ]
206   #a #l1 #x0 #a0 #l2 #l3 #Hls #Hrs #Hl1marks >Hls in Ht3; >Hrs #Ht3
207   cases (H2 … Ht3 ?)
208   [ * #Hc' #Ht2 % %{false} % // <Hc' @Ht2
209   | * #Hnotfalse whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
210     [ * #Hc' #Ht2 % %{true} % // <Hc' @Ht2
211     |  * #Hnottrue whd in ⊢ (%→?); #Ht2 cases (Ht2 … Ht3 ?) -Ht2
212       [ * #Hc' #Ht2 %2 <Hc' % // @Ht2
213       | * #Hnotnull @False_ind
214         generalize in match Hnotnull;generalize in match Hnottrue;generalize in match Hnotfalse;
215         cases c in Hc; normalize
216         [ * [ #_ #_ * #Hfalse #_ | #_ * #Hfalse #_ #_ ]
217         | #_ #_ #_ * #Hfalse
218         |*: #Hfalse destruct (Hfalse) ] @Hfalse %
219       | @Hl1marks ]
220     | @Hl1marks ]
221   | @Hl1marks ]
222 | #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1
223   >Ht1 in H1; #H1 cases (H1 … (refl ??)) #_ #Ht3 cases (H1 ? (refl ??)) -H1
224   #Hc #Ht3 % //
225 ]
226 qed.
227
228 (*
229 1) il primo carattere è marcato
230 2) l'ultimo carattere è l'unico che può essere null, gli altri sono bit
231 3) il terminatore non è né bit, né null
232 *)
233    
234 definition copy0 ≝ whileTM ? copy_step (inr … (inl … (inr … start_nop))).
235
236 let rec merge_config (l1,l2:list STape) ≝ 
237   match l1 with
238   [ nil ⇒ nil ?
239   | cons p1 l1' ⇒ match l2 with
240     [ nil ⇒ nil ? 
241     | cons p2 l2' ⇒ 
242            let 〈c1,b1〉 ≝ p1 in let 〈c2,b2〉 ≝ p2 in
243            match c2 with
244            [ null ⇒ p1
245            | _ ⇒ p2 ] :: merge_config l1' l2' ] ].
246            
247 lemma merge_config_append :
248  ∀l1,l2,l3,l4.|l1| = |l2| → 
249  merge_config (l1@l3) (l2@l4) = merge_config l1 l2@merge_config l3 l4.
250 #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen)
251 [normalize //
252 | #t1 #t2 * #c1 #b1 * #c2 #b2 #IH whd in ⊢ (??%%); >IH % ]
253 qed.
254
255 definition R_copy0 ≝ λt1,t2.
256   ∀ls,c,c0,rs,l1,l3,l4.
257   t1 = midtape STape (l3@l4@〈c0,true〉::ls) 〈c,true〉 (l1@rs) → 
258   no_marks l1 → no_marks (l3@l4) → |l1| = |l4| → 
259   ∀l1',bv.〈c,false〉::l1 = l1'@[〈comma,bv〉] → only_bits_or_nulls l1' → 
260   ∀l4',bg.l4@[〈c0,false〉] = 〈grid,bg〉::l4' → only_bits_or_nulls l4' → 
261   (c = comma ∧ t2 = t1) ∨
262   (c ≠ comma ∧ 
263     t2 = midtape ? (reverse ? l1'@l3@〈grid,true〉::
264                   merge_config l4' (reverse ? l1')@ls) 
265      〈comma,true〉 rs).
266      
267 lemma inj_append_singleton_l1 :
268   ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → l1 = l2.
269 #A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
270 >reverse_append >reverse_append normalize #H1 destruct
271 lapply (eq_f … (reverse ?) … e0) >reverse_reverse >reverse_reverse //
272 qed.
273
274 lemma inj_append_singleton_l2 :
275   ∀A.∀l1,l2:list A.∀a1,a2.l1@[a1] = l2@[a2] → a1 = a2.
276 #A #l1 #l2 #a1 #a2 #H lapply (eq_f … (reverse ?) … H)
277 >reverse_append >reverse_append normalize #H1 destruct %
278 qed.
279
280 axiom daemon : ∀P:Prop.P.
281
282 lemma wsem_copy0 : WRealize ? copy0 R_copy0.
283 #intape #k #outc #Hloop 
284 lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop
285 * #ta * #Hstar @(star_ind_l ??????? Hstar)
286 [ #tb whd in ⊢ (%→?); #Hleft
287   #ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
288   #Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits
289   cases (Hleft … Htb) -Hleft #Hc #Houtc % %
290   [ generalize in match Hl1bits; -Hl1bits cases l1' in Hl1;
291     [ normalize #Hl1 #c1 destruct (Hl1) %
292     | * #c' #b' #l0 #Heq normalize in Heq; destruct (Heq)
293       #Hl1bits lapply (Hl1bits 〈c',false〉 ?) [ @memb_hd ] 
294       >Hc #Hfalse destruct ]
295   | @Houtc ]
296 | #tb #tc #td whd in ⊢ (%→?→(?→%)→%→?); #Htc #Hstar1 #Hind #Htd
297   lapply (Hind Htd) -Hind #Hind
298   #ls #c #c0 #rs #l1 #l3 #l4 #Htb #Hl1nomarks #Hl3l4nomarks #Hlen #l1' #bv
299   #Hl1 #Hl1bits #l4' #bg #Hl4 #Hl4bits %2
300   cases (Htc … Htb) -Htc #Hcbitnull #Htc
301   % [ % #Hc' >Hc' in Hcbitnull; normalize #Hfalse destruct (Hfalse) ]
302   cut (|l1| = |reverse ? l4|) [>length_reverse @Hlen] #Hlen1
303   @(list_cases2 … Hlen1)
304   [ (* case l1 = [] is discriminated because l1 contains at least comma *)
305     #Hl1nil @False_ind >Hl1nil in Hl1; cases l1' normalize
306     [ #Hl1 destruct normalize in Hcbitnull; destruct (Hcbitnull)
307     | #p0 #l0 normalize #Hfalse destruct (Hfalse) cases l0 in e0;
308       [ normalize #Hfalse1 destruct (Hfalse1)
309       | #p0' #l0' normalize #Hfalse1 destruct (Hfalse1) ] ]
310   | (* case c::l1 = c::a::l1'' *)
311     * #a #ba * #a0 #ba0 #l1'' #l4'' #Hl1cons #Hl4cons
312     lapply (eq_f ?? (reverse ?) ?? Hl4cons) >reverse_reverse >reverse_cons -Hl4cons #Hl4cons
313     cut (ba = false) 
314     [ >Hl1cons in Hl1nomarks; #Hl1nomarks lapply (Hl1nomarks 〈a,ba〉 ?)
315       [ @memb_hd | normalize // ] ] #Hba
316     cut (ba0 = false) 
317     [ >Hl4cons in Hl3l4nomarks; #Hl3l4nomarks lapply (Hl3l4nomarks 〈a0,ba0〉 ?)
318       [ @memb_append_l2 @memb_append_l2 @memb_hd | normalize // ] ] #Hba0
319     >Hba0 in Hl4cons; >Hba in Hl1cons; -Hba0 -Hba #Hl1cons #Hl4cons
320     >Hl4cons in Htc; >Hl1cons #Htc
321     lapply (Htc a (l3@reverse ? l4'') c0 a0 ls (l1''@rs) ? (refl ??) ?)
322     [ #x #Hx @Hl3l4nomarks >Hl4cons <associative_append
323       @memb_append_l1 @Hx
324     | >associative_append >associative_append %
325     | -Htc
326       cut (∃la.l1' = 〈c,false〉::la)
327       [ >Hl1cons in Hl1; cases l1'
328         [normalize #Hfalse destruct (Hfalse)
329         | #p #la normalize #Hla destruct (Hla) @(ex_intro ?? la) % ] ]
330       * #la #Hla
331       cut (∃lb.l4' = lb@[〈c0,false〉])
332       [ >Hl4cons in Hl4;
333         @(list_elim_left … l4')
334         [ #Heq lapply (eq_f … (length ?) … Heq)
335           >length_append >length_append 
336           >commutative_plus normalize >commutative_plus normalize
337           #Hfalse destruct
338         | #a1 #tl #_ #Heq 
339           >(inj_append_singleton_l2 ? (reverse ? l4''@[〈a0,false〉]) (〈grid,bg〉::tl) 〈c0,false〉 a1 Heq)
340           @ex_intro //
341       ] ] * #lb #Hlb
342       cut (|lb| = |reverse ? la|) 
343       [ >Hla in Hl1; >Hlb in Hl4; #Hl4 #Hl1
344         >(?:l1 = la@[〈comma,bv〉]) in Hlen;
345         [|normalize in Hl1; destruct (Hl1) %]
346         >(?:l4 = 〈grid,bg〉::lb)
347         [|@(inj_append_singleton_l1 ?? (〈grid,bg〉::lb) ?? Hl4) ]
348         >length_append >commutative_plus >length_reverse
349         normalize #Hlalb destruct (Hlalb) //
350       ] #Hlen2
351       *
352       (* by hyp on the first iteration step, 
353          we consider whether c = bit x or c = null *)
354       (* c = bit x *)
355       [ * #x * #Hx #Htc 
356         lapply (Hind (〈bit x,false〉::ls) a a0 rs l1'' 
357                 (〈bit x,false〉::l3) (reverse ? l4'') ????) 
358         [ >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus 
359           normalize #Hlen destruct (Hlen) //
360         | #x0 #Hx0 cases (orb_true_l … Hx0)
361           [ #Hx0eq >(\P Hx0eq) %
362           | -Hx0 #Hx0 @Hl3l4nomarks >Hl4cons
363             <associative_append @memb_append_l1 // ]
364         | #x0 #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
365         | >Htc >associative_append % 
366         | -Hind 
367           <Hl1cons <Hl4cons #Hind lapply (Hind la bv ?? lb bg ??)
368           [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
369           | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
370             @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
371           | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
372           | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
373             destruct (Hl1) // ] -Hind
374           (* by IH, we proceed by cases, whether a = comma 
375              (consequently several lists = []) or not *)          
376           *
377           [ * #Ha #Houtc1
378            cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = [])
379            [ @daemon ] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
380            >Hl1cons in Hl1; >Hla
381            >Houtc1 >Htc #Hl1
382            >Hl4cons in Hl4; >Hlb #Hl4
383            >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hx
384            cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
385            normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
386           | * #Ha #Houtc1 >Houtc1 @eq_f3 //
387             >Hla >reverse_cons >associative_append @eq_f
388             >Hx whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
389             >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %            
390           ]
391        ]
392     | (* c = null *)
393       * #Hc #Htc 
394       lapply (Hind (〈c0,false〉::ls) a a0 rs l1'' (〈null,false〉::l3) (reverse ? l4'') ????)
395       [  >Hl1cons in Hlen; >Hl4cons >length_append >commutative_plus normalize
396          #Hlen destruct (Hlen) @e0
397       | #x0 #Hx0 cases (memb_append STape ? [〈null,false〉] (l3@reverse ? l4'') … Hx0) -Hx0 #Hx0
398         [ >(memb_single … Hx0) %
399         | @Hl3l4nomarks cases (memb_append … Hx0) -Hx0 #Hx0
400           [ @memb_append_l1 //
401           | @memb_append_l2 >Hl4cons @memb_append_l1 // ]
402         ]
403       | >Hl1cons #x' #Hx0 @Hl1nomarks >Hl1cons @memb_cons //
404       | >Htc @eq_f3 // >associative_append % ] -Hind <Hl1cons <Hl4cons #Hind
405         lapply (Hind la bv ?? lb bg ??)
406           [ #x0 #Hx0 @Hl4bits >Hlb @memb_append_l1 //
407           | >Hlb in Hl4; normalize in ⊢ (%→?); #Hl4
408             @(inj_append_singleton_l1 ? l4 (〈grid,bg〉::lb) … Hl4)
409           | #x0 #Hx0 @Hl1bits >Hla @memb_cons //
410           | >Hla in Hl1; normalize in ⊢ (%→?); #Hl1
411             destruct (Hl1) // ] -Hind *
412           (* by IH, we proceed by cases, whether a = comma 
413              (consequently several lists = []) or not *)          
414           [ * #Ha #Houtc1 >Hl1cons in Hl1; >Hla
415            >Houtc1 >Htc #Hl1
416            >Hl4cons in Hl4; >Hlb #Hl4
417            cut (la = [] ∧ lb = [] ∧ l1'' = [] ∧ l4'' = []) 
418            [@daemon] * * * #Hla1 #Hlb1 #Hl1nil #Hl4nil
419            >Hla1 >Hlb1 >Hl1nil >Hl4nil >Hc
420            cut (a0 = grid) [ @daemon ] #Ha0 <Ha <Ha0
421            normalize in ⊢ (??(??%?%)(??%?%)); >associative_append %
422           | * #Ha #Houtc1 >Houtc1 @eq_f3 //
423             >Hla >reverse_cons >associative_append @eq_f
424             >Hc whd in ⊢ (??%?); @eq_f whd in ⊢ (???%); @eq_f @eq_f
425             >Hlb >append_cons @eq_f2 // >(merge_config_append … Hlen2) %
426           ]
427        ]
428 ]]]
429 qed.
430
431 definition merge_char ≝ λc1,c2.
432   match c2 with
433   [ null ⇒ c1
434   | _ ⇒ c2 ].
435   
436 lemma merge_cons : 
437   ∀c1,c2,conf1,conf2.
438   merge_config (〈c1,false〉::conf1) (〈c2,false〉::conf2) = 
439     〈merge_char c1 c2,false〉::merge_config conf1 conf2.
440 #c1 #c2 #conf1 #conf2 normalize @eq_f2 //
441 cases c2 /2/
442 qed.
443
444 lemma merge_bits : ∀l1,l2.|l1| = |l2| → only_bits l2 → merge_config l1 l2 = l2.
445 #l1 #l2 #Hlen @(list_ind2 … Hlen) //
446 #tl1 #tl2 #hd1 #hd2 #IH
447 >(eq_pair_fst_snd … hd1) >(eq_pair_fst_snd … hd2) #Hbits
448 change with (cons ???) in ⊢ (??%?); @eq_f2
449 [ cases (\fst hd2) in Hbits; 
450   [ #b #_ %
451   |*: #Hfalse lapply (Hfalse … (memb_hd …)) normalize #Hfalse1 destruct (Hfalse1) ]
452 | @IH #x #Hx @Hbits @memb_cons // ]
453 qed.
454
455 lemma merge_config_c_nil : 
456   ∀c.merge_config c [] = [].
457 #c cases c normalize //
458 qed.
459
460 lemma reverse_merge_config :
461   ∀c1,c2.|c1| = |c2| → reverse ? (merge_config c1 c2) = 
462     merge_config (reverse ? c1) (reverse ? c2).        
463 #c1 #c2 <(length_reverse ? c1) <(length_reverse ? c2) #Hlen
464 <(reverse_reverse ? c1) in ⊢ (??%?); <(reverse_reverse ? c2) in ⊢ (??%?);
465 generalize in match Hlen; @(list_ind2 … Hlen) -Hlen //
466 #tl1 #tl2 #hd1 #hd2 #IH whd in ⊢ (??%%→?); #Hlen destruct (Hlen) -Hlen
467 <(length_reverse ? tl1) in e0; <(length_reverse ? tl2) #Hlen
468 >reverse_cons >reverse_cons >(merge_config_append ???? Hlen)
469 >reverse_append >(eq_pair_fst_snd ?? hd1) >(eq_pair_fst_snd ?? hd2)
470 whd in ⊢ (??%%); @eq_f2 // @IH //
471 qed.
472
473 definition copy
474 ≝ 
475   seq STape copy0 (seq ? (move_l …) (seq ? (adv_to_mark_l … (is_marked ?))
476    (seq ? (clear_mark …) (seq ? (adv_to_mark_r … (is_marked ?)) (clear_mark …))))).
477
478 (*
479    s0, s1 = caratteri di testa dello stato
480    c0 = carattere corrente del nastro oggetto
481    c1 = carattere in scrittura sul nastro oggetto
482    
483    questa dimostrazione sfrutta il fatto che 
484    merge_config (l0@[c0]) (l1@[c1]) = l1@[merge_char c0 c1] 
485    se l0 e l1 non contengono null
486 *)
487
488 definition R_copy ≝ λt1,t2.
489   ∀ls,s0,s1,c0,c1,rs,l1,l3,l4.
490   t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) → 
491   no_marks l1 → no_marks l3 → no_marks l4 → |l1| = |l4| → 
492   only_bits (l4@[〈s0,false〉]) → only_bits (〈s1,false〉::l1) → 
493   bit_or_null c0 = true → bit_or_null c1 = true → 
494   t2 = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l3@〈grid,false〉::
495                       〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
496        〈comma,false〉 rs.
497        
498 axiom sem_copy0 : Realize ? copy0 R_copy0.
499
500 definition option_cons ≝ λA.λa:option A.λl.
501   match a with
502   [ None ⇒ l
503   | Some a' ⇒ a'::l ].
504
505 lemma sem_copy : Realize ? copy R_copy.
506 #intape 
507 cases (sem_seq … (sem_copy0 …)
508         (sem_seq … (sem_move_l …)
509           (sem_seq … (sem_adv_to_mark_l … (is_marked ?))
510             (sem_seq … (sem_clear_mark …)
511               (sem_seq … (sem_adv_to_mark_r … (is_marked ?)) (sem_clear_mark …))))) intape)
512 #k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop
513 #ls #s0 #s1 #c0 #c1 #rs #l1 #l2 #l3 #Hintape #Hl1marks #Hl2marks #Hl3marks #Hlen
514 #Hbits1 #Hbits2 #Hc0bits #Hc1bits
515 cases HR -HR #ta * whd in ⊢ (%→?); #Hta 
516 cut (ta = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l2@〈grid,true〉::
517                       〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)
518        〈comma,true〉 rs)
519 [lapply (Hta ls s1 s0 rs (l1@[〈c1,false〉;〈comma,false〉]) l2 (〈grid,false〉::〈c0,false〉::l3) ?)
520   [>associative_append in ⊢ (???(????%)); normalize in ⊢ (???(??%%%)); @Hintape ]
521  -Hta #Hta cases (Hta ??? (〈s1,false〉::l1@[〈c1,false〉]) false ? ? ?? (refl ??) ?)
522   [3: #x #Hx cases (memb_append … Hx) -Hx #Hx
523     [ @Hl1marks //
524     | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | >(memb_single … Hx) % ]] 
525   |4: #x #Hx cases (memb_append … Hx) -Hx #Hx
526     [ @Hl2marks //
527     | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | cases (orb_true_l … Hx) [-Hx #Hx >(\P Hx) % | @Hl3marks ] ] ]
528   |5: >length_append normalize >Hlen >commutative_plus %
529   |6: normalize >associative_append %
530   |7: #x #Hx cases (memb_append ?? (〈s1,false〉::l1) … Hx) -Hx #Hx
531     [ whd in ⊢ (??%?); >(Hbits2 … Hx) %
532     | >(memb_single … Hx) // ]
533   |8: #x #Hx cases (memb_append … Hx) -Hx #Hx
534     [ cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) // | whd in ⊢ (??%?); >Hbits1 // @memb_append_l1 // ]
535     | >(memb_single … Hx) whd in ⊢ (??%?); >(Hbits1 〈s0,false〉) // @memb_append_l2 @memb_hd ]
536   | * #Hs1 @False_ind >Hs1 in Hbits2; #Hbits2 lapply (Hbits2 〈comma,false〉 ?) //
537     normalize #Hfalse destruct (Hfalse)
538   | * #Hs1 #Ht2 >Ht2 >reverse_cons >reverse_append >reverse_single @eq_f3 //
539     >merge_cons >merge_bits
540     [2: #x #Hx @Hbits2 cases (memb_append STape ? (reverse ? l1) ? Hx) -Hx #Hx
541       [@daemon | >(memb_single … Hx) @memb_hd ]
542     |3: >length_append >length_append >length_reverse >Hlen % ]
543     normalize >associative_append normalize >associative_append %
544   ]
545 ] -Hta #Hta * #tb * whd in ⊢ (%→?); #Htb
546 lapply (Htb … Hta) -Htb #Htb change with (midtape ????) in Htb:(???%);
547 * #tc * whd in ⊢ (%→?); #Htc 
548 cases (Htc … Htb)
549 [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
550 * #_ #Htc 
551 lapply (Htc (reverse ? l1@〈s1,false〉::l2) 〈grid,true〉 
552           (〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)???)
553 [ #x #Hx cases (memb_append … Hx) -Hx #Hx
554   [ @Hl1marks @daemon
555   | cases (orb_true_l … Hx) -Hx #Hx
556     [ >(\P Hx) % | @(Hl2marks … Hx) ] ]
557 | %
558 | whd in ⊢ (??%?); >associative_append % ] -Htc #Htc
559 * #td * whd in ⊢ (%→?); #Htd lapply (Htd … Htc) -Htd #Htd
560 * #te * whd in ⊢ (%→?); #Hte cases (Hte … Htd) -Hte -Htd
561 [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ]
562 * #_ #Hte 
563 lapply (Hte (reverse ? (reverse ? l1@〈s1,false〉::l2)@[〈c1,false〉])
564          〈comma,true〉 rs ? (refl ??) ?) -Hte
565 [ >reverse_append >reverse_cons >reverse_reverse #x #Hx
566   cases (memb_append … Hx) -Hx #Hx
567   [ cases (memb_append … Hx) -Hx #Hx
568     [ cases (memb_append … Hx) -Hx #Hx
569       [ @daemon 
570       | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
571     | @(Hl1marks … Hx) ]
572   | lapply (memb_single … Hx) -Hx #Hx >Hx % ]
573 | >reverse_append >reverse_reverse >reverse_cons
574   >associative_append >associative_append >associative_append
575   >associative_append >associative_append % ]
576 #Hte whd in ⊢ (%→?); #Houtc lapply (Houtc … Hte) -Houtc #Houtc >Houtc
577 @eq_f3 //
578 >reverse_append >reverse_append >reverse_single >reverse_cons
579 >reverse_append >reverse_append >reverse_reverse >reverse_reverse
580 >reverse_single >associative_append >associative_append %
581 qed.