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