]> matita.cs.unibo.it Git - helm.git/commitdiff
Progress
authorWilmer Ricciotti <ricciott@cs.unibo.it>
Tue, 22 May 2012 14:27:50 +0000 (14:27 +0000)
committerWilmer Ricciotti <ricciott@cs.unibo.it>
Tue, 22 May 2012 14:27:50 +0000 (14:27 +0000)
matita/matita/lib/turing/universal/move_tape.ma
matita/matita/lib/turing/universal/uni_step.ma

index 716388c86316abfe968bc6d11bc8b820f55d5a28..d8e4779bc98f54e87ee6e1fd71c489faed144bd2 100644 (file)
@@ -354,18 +354,55 @@ definition R_uni_step ≝ λt1,t2.
 definition no_nulls ≝ 
  λl:list STape.∀x.memb ? x l = true → is_null (\fst x) = false.
  
+definition current_of_alpha ≝ λc:STape.
+  match \fst c with [ null ⇒ None ? | _ ⇒ Some ? c ].
+
+definition legal_tape ≝ λls,c,rs.
+ let t ≝ mk_tape STape ls (current_of_alpha c) rs in
+ left ? t = ls ∧ right ? t = rs ∧ current ? t = current_of_alpha c.
+lemma legal_tape_cases : 
+  ∀ls,c,rs.legal_tape ls c rs → 
+  \fst c ≠ null ∨ (\fst c = null ∧ (ls = [] ∨ rs = [])).
+#ls #c #rs cases c #c0 #bc0 cases c0
+[ #c1 normalize #_ % % #Hfalse destruct (Hfalse)
+| cases ls
+  [ #_ %2 % // % %
+  | #l0 #ls0 cases rs
+    [ #_ %2 % // %2 %
+    | #r0 #rs0 normalize * * #_ #Hrs destruct (Hrs) ]
+  ]
+|*: #_ % % #Hfalse destruct (Hfalse) ]
+qed.
+
+axiom legal_tape_conditions : 
+  ∀ls,c,rs.(\fst c ≠ null ∨ ls = [] ∨ rs = []) → legal_tape ls c rs.
+(*#ls #c #rs *
+[ *
+  [ >(eq_pair_fst_snd ?? c) cases (\fst c)
+    [ #c0 #Hc % % %
+    | * #Hfalse @False_ind /2/
+    |*: #Hc % % %
+    ]
+  | cases ls [ * #Hfalse @False_ind /2/ ]
+    #l0 #ls0 
+  
+  #Hc
+*)
 definition R_move_tape_r_abstract ≝ λt1,t2.
   ∀rs,n,table,curc,curconfig,ls.
   bit_or_null curc = true → only_bits_or_nulls curconfig → table_TM n (reverse ? table) → 
   t1 = midtape STape (table@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls) 
          〈grid,false〉 rs →
   no_nulls rs → no_marks rs → 
+  legal_tape ls 〈curc,false〉 rs → 
   ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs → 
   ∃ls1,rs1,newc.
   (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
     〈grid,false〉::reverse ? table@〈grid,false〉::rs1) ∧
    lift_tape ls1 〈newc,false〉 rs1 = 
-   tape_move_right STape ls 〈curc,false〉 rs).
+   tape_move_right STape ls 〈curc,false〉 rs ∧ legal_tape ls1 〈newc,false〉 rs1).
    
 lemma lift_tape_not_null :
   ∀ls,c,rs. is_null (\fst c) = false → 
@@ -379,22 +416,30 @@ lemma mtr_concrete_to_abstract :
   ∀t1,t2.R_move_tape_r t1 t2 → R_move_tape_r_abstract t1 t2.
 #t1 #t2 whd in ⊢(%→?); #Hconcrete
 #rs #n #table #curc #curconfig #ls #Hcurc #Hcurconfig #Htable #Ht1
-#Hrsnonulls #Hrsnomarks #t1' #Ht1'
+#Hrsnonulls #Hrsnomarks #Htape #t1' #Ht1'
 cases (Hconcrete … Htable Ht1) //
 [ * #Hrs #Ht2
   @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? [])
   @(ex_intro ?? null) %
-  [ >Ht2 %
-  | >Hrs % ]
+  [ %
+    [ >Ht2 %
+    | >Hrs % ]
+  | % % % ]
 | * * #r0 #br0 * #rs0 * #Hrs 
   cut (br0 = false) [@(Hrsnomarks 〈r0,br0〉) >Hrs @memb_hd]
   #Hbr0 >Hbr0 in Hrs; #Hrs #Ht2
   @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? rs0)
   @(ex_intro ?? r0) %
-  [ >Ht2  //
-  | >Hrs >lift_tape_not_null
-    [ %
-    | @(Hrsnonulls 〈r0,false〉) >Hrs @memb_hd ] ]
+  [ %
+    [ >Ht2  //
+    | >Hrs >lift_tape_not_null
+      [ %
+      | @(Hrsnonulls 〈r0,false〉) >Hrs @memb_hd ] ]
+  | @legal_tape_conditions % % % #Hr0 >Hr0 in Hrs; #Hrs
+    lapply (Hrsnonulls 〈null,false〉 ?)
+    [ >Hrs @memb_hd | normalize #H destruct (H) ]
+  ]
+]
 qed.
 
 definition R_move_tape_l_abstract ≝ λt1,t2.
@@ -403,35 +448,43 @@ definition R_move_tape_l_abstract ≝ λt1,t2.
   t1 = midtape STape (table@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls) 
          〈grid,false〉 rs →
   no_nulls ls → no_marks ls → 
+  legal_tape ls 〈curc,false〉 rs → 
   ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs → 
   ∃ls1,rs1,newc.
   (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
     〈grid,false〉::reverse ? table@〈grid,false〉::rs1) ∧
    lift_tape ls1 〈newc,false〉 rs1 = 
-   tape_move_left STape ls 〈curc,false〉 rs).
+   tape_move_left STape ls 〈curc,false〉 rs ∧ legal_tape ls1 〈newc,false〉 rs1).
 
 lemma mtl_concrete_to_abstract :
   ∀t1,t2.R_move_tape_l t1 t2 → R_move_tape_l_abstract t1 t2.
 #t1 #t2 whd in ⊢(%→?); #Hconcrete
 #rs #n #table #curc #curconfig #ls #Hcurc #Hcurconfig #Htable #Ht1
-#Hlsnonulls #Hlsnomarks #t1' #Ht1'
+#Hlsnonulls #Hlsnomarks #Htape #t1' #Ht1'
 cases (Hconcrete … Htable Ht1) //
 [ * #Hls #Ht2
   @(ex_intro ?? [])
   @(ex_intro ?? (〈curc,false〉::rs)) 
   @(ex_intro ?? null) %
-  [ >Ht2 %
-  | >Hls % ]
+  [ %
+    [ >Ht2 %
+    | >Hls % ]
+  | % % % ]
 | * * #l0 #bl0 * #ls0 * #Hls
   cut (bl0 = false) [@(Hlsnomarks 〈l0,bl0〉) >Hls @memb_hd]
   #Hbl0 >Hbl0 in Hls; #Hls #Ht2
   @(ex_intro ?? ls0)
   @(ex_intro ?? (〈curc,false〉::rs)) 
   @(ex_intro ?? l0) %
-  [ >Ht2 %
-  | >Hls >lift_tape_not_null
-    [ %
-    | @(Hlsnonulls 〈l0,false〉) >Hls @memb_hd ] ]
+  [ % 
+    [ >Ht2 %
+    | >Hls >lift_tape_not_null
+      [ %
+      | @(Hlsnonulls 〈l0,false〉) >Hls @memb_hd ] ]
+  | @legal_tape_conditions % % % #Hl0 >Hl0 in Hls; #Hls
+    lapply (Hlsnonulls 〈null,false〉 ?)
+    [ >Hls @memb_hd | normalize #H destruct (H) ]
+  ]
 qed.
 
 lemma Realize_to_Realize : 
@@ -504,8 +557,10 @@ definition R_move_tape ≝ λt1,t2.
   t1 = midtape STape (table1@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls) 
          〈c,false〉 (table2@〈grid,false〉::rs) →
   no_nulls ls → no_nulls rs → no_marks ls → no_marks rs → 
+  legal_tape ls 〈curc,false〉 rs → 
   ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs → 
   ∃ls1,rs1,newc.
+  legal_tape ls1 〈newc,false〉 rs1 ∧
   (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@〈newc,false〉::
     〈grid,false〉::reverse ? table1@〈c,false〉::table2@〈grid,false〉::rs1) ∧
    ((c = bit false ∧ lift_tape ls1 〈newc,false〉 rs1 = tape_move_left STape ls 〈curc,false〉 rs) ∨
@@ -523,7 +578,7 @@ cases (sem_if ? (test_char ??) … tc_true (sem_test_char ? (λc:STape.c == 〈b
 #k * #outc * #Hloop #HR
 @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
 #rs #n #table1 #c #table2 #curc #curconfig #ls
-#Hcurc #Hc #Hcurconfig #Htable #Hintape #Hls #Hrs #Hls1 #Hrs1 #t1' #Ht1'
+#Hcurc #Hc #Hcurconfig #Htable #Hintape #Hls #Hrs #Hls1 #Hrs1 #Htape #t1' #Ht1'
 generalize in match HR; -HR *
 [ * #ta * whd in ⊢ (%→?); #Hta cases (Hta 〈c,false〉 ?)
   [| >Hintape % ] -Hta #Hceq #Hta lapply (\P Hceq) -Hceq #Hceq destruct (Hta Hceq)
@@ -533,11 +588,14 @@ generalize in match HR; -HR *
   [ @daemon ] -Htb >append_cons <associative_append #Htb
   whd in ⊢ (%→?); #Houtc lapply (Houtc … Htb … Ht1') //
   [ >reverse_append >reverse_append >reverse_reverse @Htable |]
-  -Houtc -Htb * #ls1 * #rs1 * #newc * #Houtc #Hnewtape
+  -Houtc -Htb * #ls1 * #rs1 * #newc * * #Houtc #Hnewtape #Hnewtapelegal
   @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) % 
-  [ >Houtc >reverse_append >reverse_append >reverse_reverse
-    >associative_append >associative_append % 
-  | % % % // ]
+  [ //
+  | % 
+    [ >Houtc >reverse_append >reverse_append >reverse_reverse
+      >associative_append >associative_append % 
+    | % % % // ]
+  ]
 | * #ta * whd in ⊢ (%→?); #Hta cases (Hta 〈c,false〉 ?) 
   [| >Hintape % ] -Hta #Hcneq cut (c ≠ bit false) 
   [ lapply (\Pf Hcneq) @not_to_not #Heq >Heq % ] -Hcneq #Hcneq #Hta destruct (Hta)
@@ -550,11 +608,14 @@ generalize in match HR; -HR *
     [ @daemon ] -Htc >append_cons <associative_append #Htc
     whd in ⊢ (%→?); #Houtc lapply (Houtc … Htc … Ht1') //
     [ >reverse_append >reverse_append >reverse_reverse @Htable |]
-    -Houtc -Htc * #ls1 * #rs1 * #newc * #Houtc #Hnewtape
+    -Houtc -Htc * #ls1 * #rs1 * #newc * * #Houtc #Hnewtape #Hnewtapelegal
     @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) % 
-    [ >Houtc >reverse_append >reverse_append >reverse_reverse
-      >associative_append >associative_append % 
-    | % %2 % // ]
+    [ //
+    | %
+      [ >Houtc >reverse_append >reverse_append >reverse_reverse
+        >associative_append >associative_append % 
+      | % %2 % // ]
+    ]
   | * #tb * whd in ⊢ (%→?); #Htb cases (Htb 〈c,false〉 ?) 
     [| >Hintape % ] -Htb #Hcneq' cut (c ≠ bit true) 
     [ lapply (\Pf Hcneq') @not_to_not #Heq >Heq % ] -Hcneq' #Hcneq' #Htb destruct (Htb)
@@ -566,12 +627,15 @@ generalize in match HR; -HR *
     [ * >(bit_or_null_not_grid … Hcurc) #Hfalse destruct (Hfalse) ]
     * #_ #Houtc lapply (Houtc … (refl ??) (refl ??) ?) [@daemon] -Houtc #Houtc
     @(ex_intro ?? ls) @(ex_intro ?? rs) @(ex_intro ?? curc) %
-    [ @Houtc
-    | %2 % // % // % // 
-      generalize in match Hcneq; generalize in match Hcneq'; 
-      cases c in Hc; normalize //
-      [ * #_ normalize [ #Hfalse @False_ind cases Hfalse /2/ | #_ #Hfalse @False_ind cases Hfalse /2/ ] 
-      |*: #Hfalse destruct (Hfalse) ]
+    [ //
+    | %
+      [ @Houtc
+      | %2 % // % // % // 
+        generalize in match Hcneq; generalize in match Hcneq'; 
+        cases c in Hc; normalize //
+        [ * #_ normalize [ #Hfalse @False_ind cases Hfalse /2/ | #_ #Hfalse @False_ind cases Hfalse /2/ ] 
+        |*: #Hfalse destruct (Hfalse) ]
+      ]
     ]
   ]
 ]
index 6a6bdc07cd297b395b80e9a796c71c7846f7e8c3..0cc8f71c49140c35a2d2ef564c3302f34c3a9057 100644 (file)
@@ -244,31 +244,18 @@ definition exec_move ≝
 definition map_move ≝ 
   λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
 
-definition current_of_alpha ≝ λc:STape.
-  match \fst c with [ null ⇒ None ? | _ ⇒ Some ? c ].
-
-definition legal_tape ≝ λls,c,rs.
- let t ≝ mk_tape STape ls (current_of_alpha c) rs in
- left ? t = ls ∧ right ? t = rs ∧ current ? t = current_of_alpha c.
-lemma legal_tape_cases : 
-  ∀ls,c,rs.legal_tape ls c rs → 
-  \fst c ≠ null ∨ (\fst c = null ∧ ls = []) ∨ (\fst c = null ∧ rs = []).
-#ls #c #rs cases c #c0 #bc0 cases c0
-[ #c1 normalize #_ % % % #Hfalse destruct (Hfalse)
-| cases ls
-  [ #_ % %2 % %
-  | #l0 #ls0 cases rs
-    [ #_ %2 % %
-    | #r0 #rs0 normalize * * #Hls #Hrs destruct (Hrs) ]
-  ]
-|*: #_ % % % #Hfalse destruct (Hfalse) ]
-qed.
-
+(* - aggiungere a legal_tape le condizioni
+       only_bits ls, rs; bit_or_null c
+   - ci vuole un lemma che dimostri 
+       bit_or_null c1 = true    bit_or_null mv = true
+       mv ≠ null → c1 ≠ null
+     dal fatto che c1 e mv sono contenuti nella table
+ *)
 definition R_exec_move ≝ λt1,t2.
   ∀n,curconfig,ls,rs,c0,c1,s0,s1,table1,newconfig,mv,table2.
   table_TM n (table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2) → 
   no_marks curconfig → only_bits (curconfig@[〈s0,false〉]) → only_bits (〈s1,false〉::newconfig) → 
+  |curconfig| = |newconfig| → 
   no_nulls ls → no_nulls rs → no_marks ls → no_marks rs → 
   legal_tape ls 〈c0,false〉 rs →
   t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉 
@@ -279,7 +266,7 @@ definition R_exec_move ≝ λt1,t2.
     (〈s1,false〉::newconfig@〈c2,false〉::〈grid,false〉::
      table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs1) ∧   
   lift_tape ls1 〈c2,false〉 rs1 = 
-  tape_move STape t1' (map_move c1 mv).
+  tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1.
   
 (* move the following 2 lemmata to mono.ma *)
 lemma tape_move_left_eq :
@@ -296,6 +283,21 @@ lemma tape_move_right_eq :
 //
 qed.
 
+lemma lift_tape_not_null : 
+ ∀ls,c,bc,rs.c ≠ null → lift_tape ls 〈c,bc〉 rs = midtape ? ls 〈c,bc〉 rs.
+#ls #c #bc #rs cases c //
+#Hfalse @False_ind /2/
+qed.
+
+lemma merge_char_not_null :
+  ∀c1,c2.c1 ≠ null → merge_char c1 c2 ≠ null.
+#c1 #c2 @not_to_not cases c2
+[ #c1' normalize #Hfalse destruct (Hfalse)
+| normalize //
+| *: normalize #Hfalse destruct (Hfalse)
+]
+qed.
+
 lemma sem_exec_move : Realize ? exec_move R_exec_move.
 #intape
 cases (sem_seq … sem_init_copy
@@ -304,24 +306,33 @@ cases (sem_seq … sem_init_copy
 #k * #outc * #Hloop #HR
 @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop
 #n #curconfig #ls #rs #c0 #c1 #s0 #s1 #table1 #newconfig #mv #table2
-#Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hls #Hrs #Hls1 #Hrs1 #Htape #Hintape #t1' #Ht1'
+#Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hlen #Hls #Hrs #Hls1 #Hrs1 #Htape #Hintape #t1' #Ht1'
 cases HR -HR #ta * whd in ⊢ (%→?); #Hta
 lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1
         (newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) … Hintape) -Hta
 [ (*Hcurconfig2*) @daemon
 | (*Htable*) @daemon
-| (*FIXME*) @daemon
-| (*FIXME*) @daemon
+| (*bit_or_null c0 = true *) @daemon
+| (*Hcurconfig1*) @daemon
 | #Hta * #tb * whd in ⊢ (%→?); #Htb
   lapply (Htb (〈grid,false〉::ls) s0 s1 c0 c1 (〈mv,false〉::table2@〈grid,false〉::rs) newconfig (〈comma,false〉::reverse ? table1) curconfig Hta ????????) -Htb
-  [9:|*:(* rivedere le precondizioni *) @daemon ]
+  [9:|*:(* bit_or_null c0,c1; |curconfig| = |newconfig|*) @daemon ]
   #Htb * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc whd in ⊢(???(??%%%)→?);#Htc
   whd in ⊢ (%→?); #Houtc whd in Htc:(???%); whd in Htc:(???(??%%%));
   lapply (Houtc rs n 
     (〈comma,false〉::〈c1,false〉::reverse ? newconfig@〈s1,false〉::〈comma,false〉::reverse ? table1)
     mv table2 (merge_char c0 c1) (reverse ? newconfig@[〈s1,false〉]) ls ????? 
-    Hls Hrs Hls1 Hrs1 ??)
-  [3: >Htc @(eq_f3 … (midtape ?))
+    Hls Hrs Hls1 Hrs1 ???)
+  [3: cases (true_or_false (c0 ==  null)) #Hc0
+    [ >(\P Hc0) in Htape; #Htape cases (legal_tape_cases … Htape)
+      [ * #Hfalse @False_ind /2/
+      | * #_ *
+        [ #Hlsnil @legal_tape_conditions % %2 //
+        | #Hrsnil @legal_tape_conditions %2 //
+        ]
+      ]
+    | @legal_tape_conditions % % @merge_char_not_null @(\Pf Hc0) ]
+  |4:>Htc @(eq_f3 … (midtape ?))
     [ @eq_f @eq_f >associative_append >associative_append %
     | %
     | % ]
@@ -331,70 +342,84 @@ lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1
      >associative_append >associative_append >associative_append
      >associative_append >associative_append
      @Htable
-  | (* only bits or nulls c1,c2 → only bits or nulls (merge c1 c2) *) @daemon
-  | (* add to hyps? *) @daemon
+  | (* Hnewconfig *) @daemon
+  | (* bit_or_null mv = true *) @daemon
   | (* bit_or_null c1,c2 → bit_or_null (merge_char c1 c2) *) @daemon
-  | -Houtc * #ls1 * #rs1 * #newc * #Houtc *
+  | -Houtc * #ls1 * #rs1 * #newc * #Hnewtapelegal * #Houtc *
     [ *
       [ * #Hmv #Htapemove
         @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc)
         %
-        [ >Houtc -Houtc >reverse_append
-          >reverse_reverse >reverse_single @eq_f
-          >reverse_cons >reverse_cons >reverse_append >reverse_cons
-          >reverse_cons >reverse_reverse >reverse_reverse
-          >associative_append >associative_append
-          >associative_append >associative_append
-          >associative_append >associative_append %
-        | >Hmv >Ht1' >Htapemove 
-          (* mv = bit false -→ c1 = bit ?        
-          whd in Htapemove:(???%); whd in ⊢ (???%);
-          whd in match (lift_tape ???) in ⊢ (???%);
-          >Htapemove *)
-          cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
-          >Hc1 >tape_move_left_eq cases c0
-          [ #c0' %
-          | cases ls
-            [ cases rs 
-              [ %
-              | #r0 #rs0 % ]
-            | #l0 #ls0 cases rs
-              [ %
-              | #r0 #rs0 
-              
-              ls #... null # ... # rs
-              ls #... c # ... # rs
-              
-              ∃t.left ? t = ls, right ? t = rs, current t = [[ c ]]
-              
-              % ]
-            
-          
-           @eq_f3
-          [
-          
-           whd in match (merge_char ??);
-          whd in match (map_move ??);
-          
+        [ %
+          [ >Houtc -Houtc >reverse_append
+            >reverse_reverse >reverse_single @eq_f
+            >reverse_cons >reverse_cons >reverse_append >reverse_cons
+            >reverse_cons >reverse_reverse >reverse_reverse
+            >associative_append >associative_append
+            >associative_append >associative_append
+            >associative_append >associative_append %
+          | >Hmv >Ht1' >Htapemove 
+            (* mv = bit false -→ c1 = bit ? *)
+            cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
+            >Hc1 >tape_move_left_eq cases (legal_tape_cases … Htape)
+            [ #Hc0 >lift_tape_not_null /2/
+            | * #Hc0 *
+              [ #Hls >Hc0 >Hls whd in ⊢ (??%%); cases rs; 
+                [%| #r0 #rs0 %]
+              | #Hrs >Hc0 >Hrs cases ls; [%| #l0 #ls0 %]
+              ]
+            ]
+          ]
+        | //
         ]
-      |
-    
-    
-  
-   >Heqcurconfig
->append_cons in Hintape; >associative_append
-cases (Hta 
-
- cases (Hta … Hintape) -Hta
-[ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
-* #_ #Hta lapply (Hta ? 〈comma,true〉 … (refl ??) (refl ??) ?)
-[ @daemon ] -Hta #Hta
-* #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta)
-
-definition move_of_unialpha ≝ 
-  λc.match c with
-  [ bit x ⇒ match x with [ true ⇒ R | false ⇒ L ]
-  | _ ⇒ N ].
+      | * #Hmv #Htapemove 
+        @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
+        [ %
+          [ >Houtc -Houtc >reverse_append
+            >reverse_reverse >reverse_single @eq_f
+            >reverse_cons >reverse_cons >reverse_append >reverse_cons
+            >reverse_cons >reverse_reverse >reverse_reverse
+            >associative_append >associative_append
+            >associative_append >associative_append
+            >associative_append >associative_append %
+          |>Hmv >Ht1' >Htapemove 
+            cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
+            >Hc1 >tape_move_right_eq cases (legal_tape_cases … Htape)
+            [ #Hc0 >lift_tape_not_null /2/
+            | * #Hc0 *
+              [ #Hls >Hc0 >Hls whd in ⊢ (??%%); cases rs; 
+                [%| #r0 #rs0 %]
+              | #Hrs >Hc0 >Hrs cases ls; [%| #l0 #ls0 %]
+            ]
+          ]
+        ]
+      | //
+      ]
+    ]
+  | * * * #Hmv #Hlseq #Hrseq #Hnewc 
+    @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) %
+    [ %
+      [ >Houtc -Houtc >reverse_append
+        >reverse_reverse >reverse_single @eq_f
+        >reverse_cons >reverse_cons >reverse_append >reverse_cons
+        >reverse_cons >reverse_reverse >reverse_reverse
+        >associative_append >associative_append
+        >associative_append >associative_append
+        >associative_append >associative_append %
+      |>Hmv >Ht1' cases c1 in Hnewc;
+        [ #c1' whd in ⊢ (??%?→?);#Hnewc <Hnewc
+          >Hlseq >Hrseq cases Htape * #Hleft #Hright #_
+          whd in ⊢ (??%%); >Hleft >Hright % 
+        | whd in ⊢ (??%?→?); #Hnewc >Hnewc >Hlseq >Hrseq %
+        |*: whd in ⊢ (??%?→?);#Hnewc <Hnewc
+         >Hlseq >Hrseq cases Htape * #Hleft #Hright #_
+         whd in ⊢ (??%%); >Hleft >Hright % ]
+      ]
+    | //
+    ]
+  ]
+]
+qed.
 
 definition R_uni_step ≝ λt1,t2.
   ∀n,table,c,c1,ls,rs,curs,curc,news,newc,mv.