]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/universal/uni_step.ma
Progress
[helm.git] / matita / matita / lib / turing / universal / uni_step.ma
index c22628d1b0f19bd979a11cc826f768d7d5ea536f..128c81aa57da708ea32bf4b5298954ce1e799daf 100644 (file)
@@ -94,14 +94,77 @@ whd in ⊢ (%→?); #Htf cases (Htf … Hte) -Htf -Hte
   [#x #membl @Hnomarks @daemon] -Htf #Htf >Htf >reverse_reverse %
 qed.
 
+
 (* init_copy 
 
-           adv_mark_r;
            init_current_on_match; (* no marks in current *)
            move_r;
            adv_to_mark_r;
+           adv_mark_r;
+
 *)
-     
+
+definition init_copy ≝ 
+  seq ? init_current_on_match
+    (seq ? (move_r ?) 
+      (seq ? (adv_to_mark_r ? (is_marked ?))
+        (adv_mark_r ?))).
+
+definition R_init_copy ≝ λt1,t2.
+  ∀l1,l2,c,ls,d,rs. 
+  no_marks l1 → no_grids l1 → 
+  no_marks l2 → is_grid c = false → 
+  t1 = midtape STape (l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉 (l2@〈comma,true〉::〈d,false〉::rs) → 
+  t2 = midtape STape (〈comma,false〉::(reverse ? l2)@〈grid,false〉::l1@〈c,true〉::〈grid,false〉::ls) 〈d,true〉 rs.
+
+lemma list_last: ∀A.∀l:list A.
+  l = [ ] ∨ ∃a,l1. l = l1@[a].
+#A #l <(reverse_reverse ? l) cases (reverse A l)
+  [%1 //
+  |#a #l1 %2 @(ex_intro ?? a) @(ex_intro ?? (reverse ? l1)) //
+  ]
+qed.
+   
+lemma sem_init_copy : Realize ? init_copy R_init_copy.
+#intape 
+cases (sem_seq ????? sem_init_current_on_match
+        (sem_seq ????? (sem_move_r ?)
+          (sem_seq ????? (sem_adv_to_mark_r ? (is_marked ?))
+            (sem_adv_mark_r ?))) intape)
+#k * #outc * #Hloop #HR 
+@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
+#l1 #l2 #c #ls #d #rs #Hl1marks #Hl1grids #Hl2marks #Hc #Hintape
+cases HR -HR
+#ta * whd in ⊢ (%→?); #Hta lapply (Hta … Hl1grids Hc Hintape) -Hta -Hintape #Hta
+* #tb * whd in ⊢ (%→?); #Htb lapply (Htb  … Hta) -Htb -Hta
+generalize in match Hl1marks; -Hl1marks cases (list_last ? l1) 
+  [#eql1 >eql1 #Hl1marks whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?); #Htb
+   * #tc * whd in ⊢ (%→?); #Htc lapply (Htc  … Htb) -Htc -Htb *
+    [* whd in ⊢ ((??%?)→?); #Htemp destruct (Htemp)]
+   * #_ #Htc lapply (Htc … (refl …) (refl …) ?)
+    [#x #membx @Hl2marks @membx]
+   #Htc whd in ⊢ (%→?); #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc
+   >Houtc %
+  |* #c1 * #tl #eql1 >eql1 #Hl1marks >reverse_append >reverse_single 
+   whd in ⊢ ((???%)→?); whd in ⊢ ((???(????%))→?);
+   >associative_append whd in ⊢ ((???(????%))→?); #Htb
+   * #tc * whd in ⊢ (%→?); #Htc lapply (Htc  … Htb) -Htc -Htb *
+    [* >Hl1marks [#Htemp destruct (Htemp)] @memb_append_l2 @memb_hd]
+   * #_ >append_cons <associative_append #Htc lapply (Htc … (refl …) (refl …) ?)
+    [#x #membx cases (memb_append … membx) -membx #membx
+      [cases (memb_append … membx) -membx #membx
+        [@Hl1marks @memb_append_l1 @daemon
+        |>(memb_single … membx) %
+        ]
+      |@Hl2marks @membx
+      ]]
+  #Htc whd in ⊢ (%→?); #Houtc lapply (Houtc … Htc) -Houtc -Htc #Houtc
+  >Houtc >reverse_append >reverse_append >reverse_single 
+  >reverse_reverse >associative_append >associative_append 
+  >associative_append %
+qed.
+  
+(* OLD 
 definition init_copy ≝ 
   seq ? (adv_mark_r ?) 
     (seq ? init_current_on_match
@@ -169,136 +232,343 @@ cases HR -HR
       >associative_append % 
     ]
   ]
-qed.
+qed. *)
 
 include "turing/universal/move_tape.ma".
 
 definition exec_move ≝ 
-  seq ? (adv_to_mark_r … (is_marked ?))
-    (seq ? init_copy
-      (seq ? copy
-        (seq ? (move_r …)
-          (seq ? move_tape (move_r …))))).
-
-definition lift_tape ≝ λls,c,rs.
-  let 〈c0,b〉 ≝ c in
-  let c' ≝ match c0 with
-  [ null ⇒ None ?
-  | _ ⇒ Some ? c ]
-  in
-  mk_tape STape ls c' rs.
-  
-definition sim_current_of_tape ≝ λt.
-  match current STape t with
-  [ None ⇒ 〈null,false〉
-  | Some c0 ⇒ c0 ].
+  seq ? init_copy
+    (seq ? copy
+      (seq ? (move_r …) move_tape)).
 
-(*
- t1 =  ls # cs c # table # rs  
- let simt ≝ lift_tape ls c rs in
- let simt' ≝ move_left simt' in
- t2 = left simt'# cs (sim_current_of_tape simt') # table # right simt'
-*)
-          
-(*
-definition R_move
+definition map_move ≝ 
+  λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
 
+(* - 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.
-  ∀ls,current,table1,newcurrent,table2,rs.
-  t1 = midtape STape (current@〈grid,false〉::ls) 〈grid,false〉
-       (table1@〈comma,true〉::newcurrent@〈comma,false〉::move::table2@
-        〈grid,false〉::rs) → 
-  table_TM (table1@〈comma,false〉::newcurrent@〈comma,false〉::move::table2) →
-  t2 = midtape
-*)
+  ∀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) → bit_or_null c1 = true → 
+  |curconfig| = |newconfig| → 
+  legal_tape ls 〈c0,false〉 rs →
+  t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉 
+    (table1@〈comma,true〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) → 
+  ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs → 
+  ∃ls1,rs1,c2.
+  t2 = midtape STape ls1 〈grid,false〉 
+    (〈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) ∧ legal_tape ls1 〈c2,false〉 rs1.
+  
+(* move the following 2 lemmata to mono.ma *)
+lemma tape_move_left_eq :
+  ∀A.∀t:tape A.∀c.
+  tape_move ? t (Some ? 〈c,L〉) = 
+  tape_move_left ? (left ? t) c (right ? t).
+//
+qed.
 
-(*
+lemma tape_move_right_eq :
+  ∀A.∀t:tape A.∀c.
+  tape_move ? t (Some ? 〈c,R〉) = 
+  tape_move_right ? (left ? t) c (right ? t).
+//
+qed.
 
-step :
+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.
 
-if is_true(current) (* current state is final *)
-   then nop
-   else 
-   init_match;
-   match_tuple;
-   if is_marked(current) = false (* match ok *)
-      then exec_move; 
+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 merge_char_null : ∀c.merge_char null c = c.
+* //
+qed.
+
+lemma merge_char_cases : ∀c1,c2.merge_char c1 c2 = c1 ∨ merge_char c1 c2 = c2.
+#c1 *
+[ #c1' %2 %
+| % %
+| *: %2 % ]
+qed.
+
+(* lemma merge_char_c_bit :
+  ∀c1,c2.is_bit c2 = true → merge_char c1 c2 = c2.
+#c1 *
+[ #c2' #_ %
+|*: normalize #Hfalse destruct (Hfalse) ]
+qed.
+
+lemma merge_char_c_bit :
+  ∀c1,c2.is_null c2 = true → merge_char c1 c2 = c1.
+#c1 *
+[ #c2' #_ %
+|*: normalize #Hfalse destruct (Hfalse) ]
+qed.
+
+*)
+
+lemma sem_exec_move : Realize ? exec_move R_exec_move.
+#intape
+cases (sem_seq … sem_init_copy
+        (sem_seq … sem_copy
+          (sem_seq … (sem_move_r …) sem_move_tape )) intape)
+#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 #Hc1 #Hlen #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
+| (*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:|*:(* 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 ????????)
+  [3: cases Htape -Htape * * #Hnomarks #Hbits #Hc0 #Hlsrs % [ % [ %
+    [ #x #Hx cases (orb_true_l … Hx) #Hx'
+      [ >(\P Hx') %
+      | @Hnomarks @memb_cons // ]
+    | @Hbits ]
+    | cases (merge_char_cases c0 c1) #Hmerge >Hmerge // ]
+    | cases (true_or_false (c0 == null)) #Hc0'
+      [ cases Hlsrs -Hlsrs 
+        [ *
+          [ >(\P Hc0') * #Hfalse @False_ind /2/
+          | #Hlsnil % %2 // ]
+        | #Hrsnil %2 // ] 
+      | % % @merge_char_not_null @(\Pf Hc0') ] ]
+  |4:>Htc @(eq_f3 … (midtape ?))
+    [ @eq_f @eq_f >associative_append >associative_append %
+    | %
+    | % ]
+  | %
+  || >reverse_cons >reverse_cons >reverse_append >reverse_reverse 
+     >reverse_cons >reverse_cons >reverse_reverse
+     >associative_append >associative_append >associative_append
+     >associative_append >associative_append
+     @Htable
+  | (* well formedness of table *) @daemon
+  | (* Hnewconfig *) @daemon
+  | (* bit_or_null mv = true (well formedness of table) *) @daemon
+  | -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 ? *)
+            cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1
+            >Hc1 >tape_move_left_eq >(legal_tape_left … Htape) 
+            >(legal_tape_right … Htape) %
+          ]
+        | //
+        ]
+      | * #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 >(legal_tape_left … Htape) 
+            >(legal_tape_right … Htape) %
+          ]
+        | //
+        ]
+      ]
+    | * * * #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 whd in ⊢ (??%%);
+            >(legal_tape_left … Htape) >(legal_tape_right … Htape) %
+          | whd in ⊢ (??%?→?); #Hnewc >Hnewc >Hlseq >Hrseq %
+          |*: whd in ⊢ (??%?→?);#Hnewc <Hnewc
+            >Hlseq >Hrseq whd in ⊢ (??%%);
+            >(legal_tape_left … Htape) >(legal_tape_right … Htape) %
+          ]
+        ]
+      | //
+      ]
+    ]
+  ]
+]
+qed.
+
+(*
+if is_false(current) (* current state is not final *)
+   then init_match;
+    match_tuple;
+    if is_marked(current) = false (* match ok *)
+      then 
+           exec_move
+           move_r;
       else sink;
-        
+   else nop;
 *)
 
-definition mk_tuple ≝ λc,newc,mv.
-  c @ 〈comma,false〉:: newc @ 〈comma,false〉 :: [〈mv,false〉].
-
-inductive match_in_table (c,newc:list STape) (mv:unialpha) : list STape → Prop ≝ 
-| mit_hd : 
-   ∀tb.
-   match_in_table c newc mv (mk_tuple c newc mv@〈bar,false〉::tb)
-| mit_tl :
-   ∀c0,newc0,mv0,tb.
-   match_in_table c newc mv tb → 
-   match_in_table c newc mv (mk_tuple c0 newc0 mv0@〈bar,false〉::tb).
-
-definition move_of_unialpha ≝ 
-  λc.match c with
-  [ bit x ⇒ match x with [ true ⇒ R | false ⇒ L ]
-  | _ ⇒ N ].
-
-definition R_uni_step ≝ λt1,t2.
-  ∀n,table,c,c1,ls,rs,curs,curc,news,newc,mv.
-  table_TM n table → 
-  match_in_table (〈c,false〉::curs@[〈curc,false〉]) 
-    (〈c1,false〉::news@[〈newc,false〉]) mv table → 
-  t1 = midtape STape (〈grid,false〉::ls) 〈c,false〉 
-    (curs@〈curc,false〉::〈grid,false〉::table@〈grid,false〉::rs) → 
-  ∀t1',ls1,rs1.t1' = lift_tape ls 〈curc,false〉 rs → 
-  (t2 = midtape STape (〈grid,false〉::ls1) 〈c1,false〉 
-    (news@〈newc,false〉::〈grid,false〉::table@〈grid,false〉::rs1) ∧
-   lift_tape ls1 〈newc,false〉 rs1 = 
-   tape_move STape t1' (Some ? 〈〈newc,false〉,move_of_unialpha mv〉)).
-
-definition no_nulls ≝ 
- λl:list STape.∀x.memb ? x l = true → is_null (\fst x) = false.
-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 table → 
-  t1 = midtape STape (table@〈grid,false〉::〈curc,false〉::curconfig@〈grid,false〉::ls) 
-         〈grid,false〉 rs →
-  no_nulls rs → 
-  ∀t1'.t1' = lift_tape ls 〈curc,false〉 rs → 
-  ∃ls1,rs1,newc.
-  (t2 = midtape STape ls1 〈grid,false〉 (reverse ? curconfig@newc::
-    〈grid,false〉::reverse ? table@〈grid,false〉::rs1) ∧
-   lift_tape ls1 newc rs1 = 
-   tape_move_right STape ls 〈curc,false〉 rs).
+definition uni_step ≝ 
+  ifTM ? (test_char STape (λc.\fst c == bit false))
+    (single_finalTM ? (seq ? init_match
+      (seq ? match_tuple
+        (ifTM ? (test_char ? (λc.¬is_marked ? c))
+          (seq ? exec_move (move_r …))
+          (nop ?) (* sink *)
+          tc_true))))
+    (nop ?)
+    tc_true.
+
+definition R_uni_step_true ≝ λt1,t2.
+  ∀n,t0,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv.
+  table_TM (S n) (〈t0,false〉::table) → 
+  match_in_table (〈s0,false〉::curconfig@[〈c0,false〉]) 
+    (〈s1,false〉::newconfig@[〈c1,false〉]) mv (〈t0,false〉::table) → 
+  legal_tape ls 〈c0,false〉 rs → 
+  t1 = midtape STape (〈grid,false〉::ls) 〈s0,false〉 
+    (curconfig@〈c0,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs) → 
+  ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs → 
+  s0 = bit false ∧
+  ∃ls1,rs1,c2.
+  (t2 = midtape STape (〈grid,false〉::ls1) 〈s1,false〉 
+    (newconfig@〈c2,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs1) ∧
+   lift_tape ls1 〈c2,false〉 rs1 = 
+   tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1).
    
-lemma lift_tape_not_null :
-  ∀ls,c,rs. is_null (\fst c) = false → 
-  lift_tape ls c rs = mk_tape STape ls (Some ? c) rs.
-#ls * #c0 #bc0 #rs cases c0
-[|normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
-//
+definition R_uni_step_false ≝ λt1,t2.
+  ∀b. current STape t1 = Some ? 〈bit b,false〉 → b = true ∧ t2 = t1.
+  
+axiom sem_match_tuple : Realize ? match_tuple R_match_tuple.
+
+lemma sem_uni_step :
+  accRealize ? uni_step (inr … (inl … (inr … 0)))
+     R_uni_step_true R_uni_step_false. 
+@(acc_sem_if_app … (sem_test_char ? (λc:STape.\fst c == bit false))
+   (sem_seq … sem_init_match
+     (sem_seq … sem_match_tuple
+       (sem_if … (sem_test_char ? (λc.¬is_marked ? c))
+          (sem_seq … sem_exec_move (sem_move_r …))
+          (sem_nop …))))
+   (sem_nop …)
+   …)
+[ #intape #outtape 
+  #ta whd in ⊢ (%→?); #Hta #HR
+  #n #t0 #table #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
+  #Htable #Hmatch #Htape #Hintape #t1' #Ht1'
+  >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta 
+  #Hs0 lapply (\P Hs0) -Hs0 #Hs0 #Hta % //
+  cases HR -HR 
+  #tb * whd in ⊢ (%→?); #Htb 
+  lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) s0 t0 ???)
+  [ >Hta >associative_append %
+  | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @daemon
+  | (* idem *) @daemon
+  | -Hta -Htb #Htb * 
+    #tc * whd in ⊢ (%→?); #Htc cases (Htc … Htable … Htb) -Htb -Htc
+    [| * #Hcurrent #Hfalse @False_ind
+      (* absurd by Hmatch *) @daemon
+    | >Hs0 %
+    | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @daemon
+    | (* Htable *) @daemon
+    | (* Htable, Hmatch → |config| = n *) @daemon ]
+    * #table1 * #newc * #mv1 * #table2 * #Htableeq #Htc *
+    [ * #td * whd in ⊢ (%→?); >Htc -Htc #Htd
+      cases (Htd ? (refl ??)) #_ -Htd
+      cut (newc = 〈s1,false〉::newconfig@[〈c1,false〉]) [@daemon] #Hnewc
+      >Hnewc #Htd 
+      cut (mv1 = 〈\fst mv1,false〉)
+      [ >(eq_pair_fst_snd … mv1) @eq_f (*Htable, Htableeq*) @daemon ] #Hmv1
+      * #te * whd in ⊢ (%→?); #Hte
+      cut (td = midtape STape (〈c0,false〉::reverse STape curconfig@〈s0,false〉::〈grid,false〉::ls) 
+              〈grid,false〉
+              ((table1@〈s0,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉::
+               newconfig@〈c1,false〉::〈comma,false〉::〈\fst mv1,false〉::table2@〈grid,false〉::rs))
+      [ >Htd @eq_f3 //
+        [ >reverse_append >reverse_single %
+        | >associative_append >associative_append normalize
+          >associative_append >associative_append <Hmv1 %
+        ]
+      ]
+      -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte
+      [ //
+      | (*|curconfig| = |newconfig|*) @daemon
+      | (* Htable → bit_or_null c1 = true *) @daemon
+      | (* only_bits (〈s1,false〉::newconfig) *) @daemon
+      | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon
+      | (* no_marks (reverse ? curconfig) *) @daemon
+      | <Hmv1 >Hnewc in Htableeq; 
+        >associative_append >associative_append normalize
+        >associative_append >associative_append
+        #Htableeq <Htableeq // ]
+      * #ls1 * #rs1 * #c2 * * #Hte #Hliftte #Hlegalte
+      whd in ⊢ (%→?); #Houttape lapply (Houttape … Hte) -Houttape #Houttape
+      whd in Houttape:(???%); whd in Houttape:(???(??%%%));
+      @ex_intro [| @(ex_intro ?? rs1) @ex_intro [| % [ % 
+      [ >Houttape @eq_f @eq_f @eq_f @eq_f
+        change with ((〈t0,false〉::table)@?) in ⊢ (???%);
+        >Htableeq >associative_append >associative_append 
+        >associative_append normalize >associative_append
+        >associative_append normalize >Hnewc <Hmv1
+        >associative_append normalize >associative_append % 
+      | >(?: mv = \fst mv1) [| (*Hmatch, Htableeq*) @daemon ]
+        @Hliftte
+      ]
+     | //
+     ]
+    ]
+   ] 
+  | * #td * whd in ⊢ (%→%→?); >Htc #Htd
+    cases (Htd ? (refl ??)) normalize in ⊢ (%→?);
+    #Hfalse destruct (Hfalse)
+  ]
+ ]
+| #t1 #t2 #t3 whd in ⊢ (%→%→?); #Ht1 #Ht2
+  #b #Hb cases (Ht1 ? Hb) #Hb' #Ht3 >Ht2 % //
+  cases b in Hb'; normalize #H1 //
+]
 qed.
-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 #t1' #Ht1'
-cases (Hconcrete … Htable Ht1) //
-[ * #Hrs #Ht2
-  @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? [])
-  @(ex_intro ?? 〈null,false〉) %
-  [ >Ht2 %
-  | >Hrs % ]
-| * #r0 * #rs0 * #Hrs #Ht2 
-  @(ex_intro ?? (〈curc,false〉::ls)) @(ex_intro ?? rs0)
-  @(ex_intro ?? r0) %
-  [ >Ht2 %
-  | >Hrs >lift_tape_not_null
-    [ %
-    | @Hrsnonulls >Hrs @memb_hd ] ]
-qed.
\ No newline at end of file