]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/unistep_aux.ma
progress
[helm.git] / matita / matita / lib / turing / multi_universal / unistep_aux.ma
index f240cea9788d5ab6810beb3b3dbce4acb6c8af01..4be6e129f3a7f5f37d4c41312a040093f7a88780 100644 (file)
@@ -10,6 +10,9 @@
       V_____________________________________________________________*)
 
 include "turing/multi_universal/moves_2.ma".
+include "turing/multi_universal/match.ma".
+include "turing/multi_universal/copy.ma".
+include "turing/multi_universal/alphabet.ma".
 
 (*
 
@@ -48,240 +51,236 @@ include "turing/multi_universal/moves_2.ma".
   cfg_to_obj
 *)
 
-definition obj_to_cfg ≝ 
-  mmove cfg unialpha 3 L ·
-  mmove cfg unialpha 3 L ·
-  if_TM ?? (inject_TM ? (test_null ?) 3 obj)
-    (
+definition obj ≝ (0:DeqNat).
+definition cfg ≝ (1:DeqNat).
+definition prg ≝ (2:DeqNat).
 
+definition obj_to_cfg ≝
+  mmove cfg FSUnialpha 2 L ·
+  (ifTM ?? (inject_TM ? (test_null ?) 2 obj)
+    (copy_step obj cfg FSUnialpha 2 ·
+     mmove cfg FSUnialpha 2 L ·
+     mmove obj FSUnialpha 2 L) 
+    (inject_TM ? (write FSUnialpha null) 2 cfg)
+     tc_true) ·
+  inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
+  mmove cfg FSUnialpha 2 R.
+  
+definition R_obj_to_cfg ≝ λt1,t2:Vector (tape FSUnialpha) 3.
+  ∀c,ls.
+  nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls) (None ?) [ ] → 
+  (∀lso,x,rso.nth obj ? t1 (niltape ?) = midtape FSUnialpha lso x rso → 
+   t2 = change_vec ?? t1 
+         (mk_tape ? [ ] (option_hd ? (reverse ? (x::ls))) (tail ? (reverse ? (x::ls)))) cfg) ∧
+  (current ? (nth obj ? t1 (niltape ?)) = None ? → 
+   t2 = change_vec ?? t1
+         (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (null::ls))) 
+           (tail ? (reverse ? (null::ls)))) cfg).
+           
+axiom sem_move_to_end_l : ∀sig. move_to_end sig L ⊨ R_move_to_end_l sig.
+axiom accRealize_to_Realize :
+  ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
+  M ⊨ [ acc: Rtrue, Rfalse ] →  M ⊨ Rtrue ∪ Rfalse.
+  
+lemma eq_mk_tape_rightof :
+ ∀alpha,a,al.mk_tape alpha (a::al) (None ?) [ ] = rightof ? a al.
+#alpha #a #al %
+qed.
 
+axiom daemon : ∀P:Prop.P.
 
+definition option_cons ≝ λsig.λc:option sig.λl.
+  match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
 
-definition o2c_states ≝ initN 3.
-
-definition copy0 : copy_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
-definition copy1 : copy_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
-definition copy2 : copy_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
-
-
-definition trans_copy_step ≝ 
- λsrc,dst.λsig:FinSet.λn.
- λp:copy_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in
- match pi1 … q with
- [ O ⇒ match nth src ? a (None ?) with
-   [ None ⇒ 〈copy2,null_action sig n〉
-   | Some ai ⇒ match nth dst ? a (None ?) with 
-     [ None ⇒ 〈copy2,null_action ? n〉
-     | Some aj ⇒ 
-         〈copy1,change_vec ? (S n) 
-           (change_vec ? (S n) (null_action ? n) (〈None ?,R〉) src)
-           (〈Some ? ai,R〉) dst〉
-     ]
-   ]
- | S q ⇒ match q with 
-   [ O ⇒ (* 1 *) 〈copy1,null_action ? n〉
-   | S _ ⇒ (* 2 *) 〈copy2,null_action ? n〉 ] ].
-
-definition copy_step ≝ 
-  λsrc,dst,sig,n.
-  mk_mTM sig n copy_states (trans_copy_step src dst sig n) 
-    copy0 (λq.q == copy1 ∨ q == copy2).
-
-definition R_comp_step_true ≝ 
-  λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
-  ∃x,y.
-   current ? (nth src ? int (niltape ?)) = Some ? x ∧
-   current ? (nth dst ? int (niltape ?)) = Some ? y ∧
-   outt = change_vec ?? 
-            (change_vec ?? int
-              (tape_move_mono ? (nth src ? int (niltape ?)) 〈None ?, R〉) src)
-            (tape_move_mono ? (nth dst ? int (niltape ?)) 〈Some ? x, R〉) dst.
-
-definition R_comp_step_false ≝ 
-  λsrc,dst:nat.λsig,n.λint,outt: Vector (tape sig) (S n).
-    (current ? (nth src ? int (niltape ?)) = None ? ∨
-     current ? (nth dst ? int (niltape ?)) = None ?) ∧ outt = int.
-
-lemma copy_q0_q2_null :
-  ∀src,dst,sig,n,v.src < S n → dst < S n → 
-  (nth src ? (current_chars ?? v) (None ?) = None ? ∨
-   nth dst ? (current_chars ?? v) (None ?) = None ?) → 
-  step sig n (copy_step src dst sig n) (mk_mconfig ??? copy0 v) 
-  = mk_mconfig ??? copy2 v.
-#src #dst #sig #n #v #Hi #Hj
-whd in ⊢ (? → ??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (?→??%?);
-* #Hcurrent
-[ @eq_f2
-  [ whd in ⊢ (??(???%)?); >Hcurrent %
-  | whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
-| @eq_f2
-  [ whd in ⊢ (??(???%)?); >Hcurrent cases (nth src ?? (None sig)) //
-  | whd in ⊢ (??(????(???%))?); >Hcurrent
-    cases (nth src ?? (None sig)) [|#x] @tape_move_null_action ] ]
-qed.
-
-lemma copy_q0_q1 :
-  ∀src,dst,sig,n,v,a,b.src ≠ dst → src < S n → dst < S n → 
-  nth src ? (current_chars ?? v) (None ?) = Some ? a →
-  nth dst ? (current_chars ?? v) (None ?) = Some ? b → 
-  step sig n (copy_step src dst sig n) (mk_mconfig ??? copy0 v) =
-    mk_mconfig ??? copy1 
-     (change_vec ? (S n) 
-       (change_vec ?? v
-         (tape_move_mono ? (nth src ? v (niltape ?)) 〈None ?, R〉) src)
-            (tape_move_mono ? (nth dst ? v (niltape ?)) 〈Some ? a, R〉) dst).
-#src #dst #sig #n #v #a #b #Heq #Hsrc #Hdst #Ha1 #Ha2
-whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
-[ whd in match (trans ????);
-  >Ha1 >Ha2 whd in ⊢ (??(???%)?); >(\b ?) //
-| whd in match (trans ????);
-  >Ha1 >Ha2 whd in ⊢ (??(????(???%))?); >(\b ?) //
-  change with (change_vec ?????) in ⊢ (??(????%)?);
-  <(change_vec_same … v dst (niltape ?)) in ⊢ (??%?);
-  <(change_vec_same … v src (niltape ?)) in ⊢ (??%?);
-  >tape_move_multi_def 
-  >pmap_change >pmap_change <tape_move_multi_def
-  >tape_move_null_action
-  @eq_f2 // >nth_change_vec_neq //
-]
+lemma tape_move_mk_tape_R :
+  ∀sig,ls,c,rs.
+  (c = None ? → ls = [ ] ∨ rs = [ ]) → 
+  tape_move ? (mk_tape sig ls c rs) R =
+  mk_tape ? (option_cons ? c ls) (option_hd ? rs) (tail ? rs).
+#sig * [ * [ * | #c * ] | #l0 #ls0 * [ *
+[| #r0 #rs0 #H @False_ind cases (H (refl ??)) #H1 destruct (H1) ] | #c * ] ] 
+normalize //
 qed.
 
-lemma sem_copy_step :
-  ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
-  copy_step src dst sig n ⊨ 
-    [ copy1: R_comp_step_true src dst sig n, 
-            R_comp_step_false src dst sig n ].
-#src #dst #sig #n #Hneq #Hsrc #Hdst #int
-lapply (refl ? (current ? (nth src ? int (niltape ?))))
-cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
-[ #Hcur_src %{2} %
-  [| % [ %
-    [ whd in ⊢ (??%?); >copy_q0_q2_null /2/ 
-    | normalize in ⊢ (%→?); #H destruct (H) ]
-  | #_ % // % // ] ]
-| #a #Ha lapply (refl ? (current ? (nth dst ? int (niltape ?))))
-  cases (current ? (nth dst ? int (niltape ?))) in ⊢ (???%→?);
-  [ #Hcur_dst %{2} %
-    [| % [ %
-       [ whd in ⊢ (??%?); >copy_q0_q2_null /2/ 
-       | normalize in ⊢ (%→?); #H destruct (H) ]
-       | #_ % // %2 >Hcur_dst % ] ]
-  | #b #Hb %{2} %
-    [| % [ % 
-      [whd in ⊢  (??%?);  >(copy_q0_q1 … a b Hneq Hsrc Hdst) //
-      | #_ %{a} %{b} % // % //]
-      | * #H @False_ind @H %
+lemma sem_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
+@(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?)
+   (sem_seq ??????
+    (sem_if ??????????
+     (sem_test_null_multi ?? obj ?)
+      (sem_seq ?????? (accRealize_to_Realize … (sem_copy_step …))
+       (sem_seq ?????? (sem_move_multi ? 2 cfg L ?)
+        (sem_move_multi ? 2 obj L ?)))
+      (sem_inject ???? cfg ? (sem_write FSUnialpha null)))
+     (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?))
+       (sem_move_multi ? 2 cfg R ?)))) //
+#ta #tb *
+#tc * whd in ⊢ (%→?); #Htc *
+#td * *
+[ * #te * * #Hcurtc #Hte
+  * destruct (Hte) #te * *
+  [ whd in ⊢ (%→%→?); * #x * #y * * -Hcurtc #Hcurtc1 #Hcurtc2 #Hte
+    * #tf * whd in ⊢ (%→%→?); #Htf #Htd
+    * #tg * * * whd in ⊢ (%→%→%→%→?); #Htg1 #Htg2 #Htg3 #Htb
+    #c #ls #Hta1 %
+    [ #lso #x0 #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof 
+      whd in match (tape_move ???); #Htc
+      cut (tg = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg)
+      [@daemon] -Htg1 -Htg2 -Htg3 #Htg destruct (Htg Htf Hte Htd Htc Htb)
+      >change_vec_change_vec >change_vec_change_vec
+      >change_vec_commute // >change_vec_change_vec
+      >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
+      >change_vec_commute // >change_vec_change_vec
+      >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //] 
+      >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //]
+      >change_vec_commute [|@sym_not_eq //] @eq_f3 //
+      [ >Hta2 cases rso in Hta2; whd in match (tape_move_mono ???);
+        [ #Hta2 whd in match (tape_move ???); <Hta2 @change_vec_same
+        | #r1 #rs1 #Hta2 whd in match (tape_move ???); <Hta2 @change_vec_same ]
+      | >tape_move_mk_tape_R [| #_ % %] >reverse_cons
+        >nth_change_vec_neq in Hcurtc1; [|@sym_not_eq //] >Hta2
+        normalize in ⊢ (%→?); #H destruct (H) %
       ]
+    | #Hta2 >Htc in Hcurtc1; >nth_change_vec_neq [| @sym_not_eq //]
+      >Hta2 #H destruct (H)
     ]
+  | * #Hcurtc0 #Hte #_ #_ #c #ls #Hta1 >Hta1 in Htc; >eq_mk_tape_rightof
+    whd in match (tape_move ???); #Htc >Htc in Hcurtc0; *
+    [ >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //]
+      #Hcurtc #Hcurtc0 >Hcurtc0 in Hcurtc; * #H @False_ind @H %
+    | >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H) ]
   ]
+| * #te * * #Hcurtc #Hte
+  * whd in ⊢ (%→%→?); #Htd1 #Htd2
+  * #tf * * * #Htf1 #Htf2 #Htf3 whd in ⊢ (%→?); #Htb
+  #c #ls #Hta1 %
+  [ #lso #x #rso #Hta2 >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //] 
+    >Hta2 normalize in ⊢ (%→?); #H destruct (H)
+  | #_ >Hta1 in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
+    destruct (Hte) cut (td = change_vec ?? tc (midtape ? ls null []) cfg)
+    [@daemon] -Htd1 -Htd2 #Htd
+    -Htf1 cut (tf = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[null])) cfg)
+    [@daemon] -Htf2 -Htf3 #Htf destruct (Htf Htd Htc Htb)
+    >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
+    >change_vec_change_vec >change_vec_change_vec >nth_change_vec //
+    >reverse_cons >tape_move_mk_tape_R /2/ ]
 ]
 qed.
 
-definition copy ≝ λsrc,dst,sig,n.
-  whileTM … (copy_step src dst sig n) copy1.
+definition test_null_char ≝ test_char FSUnialpha (λc.c == null).
 
-definition R_copy ≝ 
-  λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
-  ((current ? (nth src ? int (niltape ?)) = None ? ∨
-    current ? (nth dst ? int (niltape ?)) = None ?) → outt = int) ∧
-  (∀ls,x,x0,rs,ls0,rs0. 
-    nth src ? int (niltape ?) = midtape sig ls x rs →
-    nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
-    (∃rs01,rs02.rs0 = rs01@rs02 ∧ |rs01| = |rs| ∧
-     outt = change_vec ?? 
-            (change_vec ?? int  
-              (mk_tape sig (reverse sig rs@x::ls) (None sig) []) src)
-            (mk_tape sig (reverse sig rs@x::ls0) (option_hd sig rs02)
-            (tail sig rs02)) dst) ∨
-    (∃rs1,rs2.rs = rs1@rs2 ∧ |rs1| = |rs0| ∧
-     outt = change_vec ?? 
-            (change_vec ?? int  
-              (mk_tape sig (reverse sig rs1@x::ls) (option_hd sig rs2)
-            (tail sig rs2)) src)
-            (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)).
+definition R_test_null_char_true ≝ λt1,t2.
+  current FSUnialpha t1 = Some ? null ∧ t1 = t2.
+  
+definition R_test_null_char_false ≝ λt1,t2.
+  current FSUnialpha t1 ≠ Some ? null ∧ t1 = t2.
+  
+lemma sem_test_null_char :
+  test_null_char ⊨ [ tc_true : R_test_null_char_true, R_test_null_char_false].
+#t1 cases (sem_test_char FSUnialpha (λc.c == null) t1) #k * #outc * * #Hloop #Htrue
+#Hfalse %{k} %{outc} % [ %
+[ @Hloop
+| #Houtc cases (Htrue ?) [| @Houtc] * #c * #Hcurt1 #Hcnull lapply (\P Hcnull)
+  -Hcnull #H destruct (H) #Houtc1 %
+  [ @Hcurt1 | <Houtc1 % ] ]
+| #Houtc cases (Hfalse ?) [| @Houtc] #Hc #Houtc %
+  [ % #Hcurt1 >Hcurt1 in Hc; #Hc lapply (Hc ? (refl ??)) 
+    >(?:((null:FSUnialpha) == null) = true) [|@(\b (refl ??)) ]
+    #H destruct (H)
+  | <Houtc % ] ]
+qed.
 
-lemma wsem_copy : ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
-  copy src dst sig n ⊫ R_copy src dst sig n.
-#src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
-lapply (sem_while … (sem_copy_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
--Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
-[ whd in ⊢ (%→?); * #Hnone #Hout %
-  [#_ @Hout
-  |#ls #x #x0 #rs #ls0 #rs0 #Hsrc1 #Hdst1 @False_ind cases Hnone
-    [>Hsrc1 normalize #H destruct (H) | >Hdst1 normalize #H destruct (H)]
-  ]
-|#tc #td * #x * #y * * #Hcx #Hcy #Htd #Hstar #IH #He lapply (IH He) -IH *
- #IH1 #IH2 %
-  [* [>Hcx #H destruct (H) | >Hcy #H destruct (H)]
-  |#ls #x' #y' #rs #ls0 #rs0 #Hnth_src #Hnth_dst
-   >Hnth_src in Hcx; whd in ⊢ (??%?→?); #H destruct (H)
-   >Hnth_dst in Hcy; whd in ⊢ (??%?→?); #H destruct (H)
-   >Hnth_src in Htd; >Hnth_dst -Hnth_src -Hnth_dst
-   cases rs
-    [(* the source tape is empty after the move *)
-     #Htd lapply (IH1 ?) 
-      [%1 >Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] >nth_change_vec //]
-     #Hout (* whd in match (tape_move ???); *) %1 %{([])} %{rs0} % 
-      [% [// | // ] 
-      |whd in match (reverse ??); whd in match (reverse ??);
-       >Hout >Htd @eq_f2 // cases rs0 //
-      ]
-    |#c1 #tl1 cases rs0
-      [(* the dst tape is empty after the move *)
-       #Htd lapply (IH1 ?) [%2 >Htd >nth_change_vec //] 
-       #Hout (* whd in match (tape_move ???); *) %2 %{[ ]} %{(c1::tl1)} % 
-        [% [// | // ] 
-        |whd in match (reverse ??); whd in match (reverse ??);
-         >Hout >Htd @eq_f2 // 
-        ]
-      |#c2 #tl2 whd in match (tape_move_mono ???); whd in match (tape_move_mono ???);
-       #Htd
-       cut (nth src (tape sig) td (niltape sig)=midtape sig (x::ls) c1 tl1)
-         [>Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] @nth_change_vec //]
-       #Hsrc_td
-       cut (nth dst (tape sig) td (niltape sig)=midtape sig (x::ls0) c2 tl2)
-         [>Htd @nth_change_vec //]
-       #Hdst_td cases (IH2 … Hsrc_td Hdst_td) -Hsrc_td -Hdst_td
-        [* #rs01 * #rs02 * * #H1 #H2 #H3 %1
-         %{(c2::rs01)} %{rs02} % [% [@eq_f //|normalize @eq_f @H2]]
-         >Htd in H3; >change_vec_commute // >change_vec_change_vec
-         >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec 
-         #H >reverse_cons >associative_append >associative_append @H 
-        |* #rs11 * #rs12 * * #H1 #H2 #H3 %2
-         %{(c1::rs11)} %{rs12} % [% [@eq_f //|normalize @eq_f @H2]]
-         >Htd in H3; >change_vec_commute // >change_vec_change_vec
-         >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec 
-         #H >reverse_cons >associative_append >associative_append @H 
-        ]
-      ]
+definition cfg_to_obj ≝
+  mmove cfg FSUnialpha 2 L ·
+  (ifTM ?? (inject_TM ? test_null_char 2 cfg)
+    (nop ? 2)
+    (copy_step cfg obj FSUnialpha 2 ·
+     mmove cfg FSUnialpha 2 L ·
+     mmove obj FSUnialpha 2 L) 
+     tc_true) ·
+  inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
+  mmove cfg FSUnialpha 2 R.
+  
+definition R_cfg_to_obj ≝ λt1,t2:Vector (tape FSUnialpha) 3.
+  ∀c,ls.
+  nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls) (None ?) [ ] → 
+  (c = null → 
+   t2 = change_vec ?? t1
+         (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (c::ls))) 
+           (tail ? (reverse ? (c::ls)))) cfg) ∧
+  (c ≠ null → 
+   t2 = change_vec ??
+          (change_vec ?? t1
+             (midtape ? (left ? (nth obj ? t1 (niltape ?))) c (right ? (nth obj ? t1 (niltape ?)))) obj)
+          (mk_tape ? [ ] (option_hd ? (reverse ? (c::ls))) (tail ? (reverse ? (c::ls)))) cfg).
+
+axiom sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
+(*@(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?)
+   (sem_seq ??????
+    (sem_if ??????????
+     (sem_test_null_multi ?? obj ?)
+      (sem_seq ?????? (accRealize_to_Realize … (sem_copy_step …))
+       (sem_move_multi ? 2 cfg L ?))
+      (sem_inject ???? cfg ? (sem_write FSUnialpha null)))
+     (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?))
+       (sem_move_multi ? 2 cfg R ?)))) //
+#ta #tb *
+#tc * whd in ⊢ (%→?); #Htc *
+#td * *
+[ * #te * * #Hcurtc #Hte
+  * destruct (Hte) #te * *
+  [ whd in ⊢ (%→%→?); * #x * #y * * -Hcurtc #Hcurtc1 #Hcurtc2 #Hte #Htd
+    * #tf * * * whd in ⊢ (%→%→%→%→?); #Htf1 #Htf2 #Htf3 #Htb
+    #c #ls #Hta1 %
+    [ #lso #x0 #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof 
+      whd in match (tape_move ???); #Htc
+      cut (tf = change_vec ?? tc (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg)
+      [@daemon] -Htf1 -Htf2 -Htf3 #Htf destruct (Htf Hte Htd Htc Htb)
+      >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
+      >nth_change_vec // >tape_move_mk_tape_R
+      @daemon
+    | #Hta2 >Htc in Hcurtc1; >nth_change_vec_neq [| @sym_not_eq //]
+      >Hta2 #H destruct (H)
     ]
+  | * #Hcurtc0 #Hte #_ #_ #c #ls #Hta1 >Hta1 in Htc; >eq_mk_tape_rightof
+    whd in match (tape_move ???); #Htc >Htc in Hcurtc0; *
+    [ >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //]
+      #Hcurtc #Hcurtc0 >Hcurtc0 in Hcurtc; * #H @False_ind @H %
+    | >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H) ]
   ]
-qed.
-     
-lemma terminate_copy :  ∀src,dst,sig,n,t.
-  src ≠ dst → src < S n → dst < S n → copy src dst sig n ↓ t.
-#src #dst #sig #n #t #Hneq #Hsrc #Hdts
-@(terminate_while … (sem_copy_step …)) //
-<(change_vec_same … t src (niltape ?))
-cases (nth src (tape sig) t (niltape ?))
-[ % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-|2,3: #a0 #al0 % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct
-| #ls #c #rs lapply c -c lapply ls -ls lapply t -t elim rs
-  [#t #ls #c % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?);
-   #H1 destruct (H1) #_ >change_vec_change_vec #Ht1 % 
-   #t2 * #x0 * #y0 * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
-   >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H)
-  |#r0 #rs0 #IH #t #ls #c % #t1 * #x * #y * * >nth_change_vec //
-   normalize in ⊢ (%→?); #H destruct (H) #Hcur
-   >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH
-  ]
+| * #te * * #Hcurtc #Hte
+  * whd in ⊢ (%→%→?); #Htd1 #Htd2
+  * #tf * * * #Htf1 #Htf2 #Htf3 whd in ⊢ (%→?); #Htb
+  #c #ls #Hta1 %
+  [ #lso #x #rso #Hta2 >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //] 
+    >Hta2 normalize in ⊢ (%→?); #H destruct (H)
+  | #_ >Hta1 in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
+    destruct (Hte) cut (td = change_vec ?? tc (midtape ? ls null []) cfg)
+    [@daemon] -Htd1 -Htd2 #Htd
+    -Htf1 cut (tf = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[null])) cfg)
+    [@daemon] -Htf2 -Htf3 #Htf destruct (Htf Htd Htc Htb)
+    >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
+    >change_vec_change_vec >change_vec_change_vec >nth_change_vec //
+    >reverse_cons >tape_move_mk_tape_R /2/ ]
 ]
 qed.
+*)
 
-lemma sem_copy : ∀src,dst,sig,n.
-  src ≠ dst → src < S n → dst < S n → 
-  copy src dst sig n ⊨ R_copy src dst sig n.
-#i #j #sig #n #Hneq #Hi #Hj @WRealize_to_Realize [/2/| @wsem_copy // ]
-qed.
+(* macchina che muove il nastro obj a destra o sinistra a seconda del valore
+   del current di prg, che codifica la direzione in cui ci muoviamo *)
+   
+definition tape_move_obj : mTM FSUnialpha 2 ≝ 
+  ifTM ?? 
+   (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit false)) 2 prg)
+   (mmove obj FSUnialpha 2 L)
+   (ifTM ?? 
+    (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit true)) 2 prg)
+    (mmove obj FSUnialpha 2 R)
+    (nop ??)
+    tc_true)
+   tc_true.
+
+definition unistep ≝ 
+  obj_to_cfg · match_m cfg prg FSUnialpha 2 · 
+  inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
+  mmove cfg FSUnialpha 2 R · copy prg cfg FSUnialpha 2 ·
+  cfg_to_obj · tape_move_obj.
\ No newline at end of file