]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/universal/uni_step.ma
progress in termination of marks.ma
[helm.git] / matita / matita / lib / turing / universal / uni_step.ma
index a2734d3330391868b8332d86de3dd25d7721f558..fc663ff2c92114aa363c1a57738b2b1696087769 100644 (file)
@@ -15,6 +15,8 @@
 *)
 
 include "turing/universal/copy.ma".
+include "turing/universal/move_tape.ma".
+include "turing/universal/match_machines.ma".
 
 (*
 
@@ -55,26 +57,23 @@ if is_true(current) (* current state is final *)
 *)
 
 definition init_match ≝ 
-  seq ? (mark ?) 
-    (seq ? (adv_to_mark_r ? (λc:STape.is_grid (\fst c)))
-      (seq ? (move_r ?) 
-        (seq ? (mark ?)
-          (seq ? (move_l ?) 
-            (adv_to_mark_l ? (is_marked ?)))))).
-            
+  mark ? · adv_to_mark_r ? (λc:STape.is_grid (\fst c)) · move_r ? · 
+    move_r ? · mark ? · move_l ? · adv_to_mark_l ? (is_marked ?).
+             
 definition R_init_match ≝ λt1,t2.
   ∀ls,l,rs,c,d. no_grids (〈c,false〉::l) → no_marks l → 
-  t1 = midtape STape ls 〈c,false〉 (l@〈grid,false〉::〈d,false〉::rs) →
-  t2 = midtape STape ls 〈c,true〉 (l@〈grid,false〉::〈d,true〉::rs).
+  t1 = midtape STape ls 〈c,false〉 (l@〈grid,false〉::〈bar,false〉::〈d,false〉::rs) →
+  t2 = midtape STape ls 〈c,true〉 (l@〈grid,false〉::〈bar,false〉::〈d,true〉::rs).
   
 lemma sem_init_match : Realize ? init_match R_init_match.
 #intape 
 cases (sem_seq ????? (sem_mark ?)
        (sem_seq ????? (sem_adv_to_mark_r ? (λc:STape.is_grid (\fst c)))
         (sem_seq ????? (sem_move_r ?)
-         (sem_seq ????? (sem_mark ?)
-          (sem_seq ????? (sem_move_l ?)
-           (sem_adv_to_mark_l ? (is_marked ?)))))) intape)
+         (sem_seq ????? (sem_move_r ?)
+          (sem_seq ????? (sem_mark ?)
+           (sem_seq ????? (sem_move_l ?)
+            (sem_adv_to_mark_l ? (is_marked ?))))))) intape)
 #k * #outc * #Hloop #HR 
 @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop
 #ls #l #rs #c #d #Hnogrids #Hnomarks #Hintape
@@ -83,18 +82,19 @@ cases HR -HR
 * #tb * whd in ⊢ (%→?); #Htb cases (Htb … Hta) -Htb -Hta 
   [* #Hgridc @False_ind @(absurd … Hgridc) @eqnot_to_noteq 
    @(Hnogrids 〈c,false〉) @memb_hd ]
-* #Hgrdic #Htb lapply (Htb l 〈grid,false〉 (〈d,false〉::rs) (refl …) (refl …) ?) 
+* #Hgrdic #Htb lapply (Htb l 〈grid,false〉 (〈bar,false〉::〈d,false〉::rs) (refl …) (refl …) ?) 
   [#x #membl @Hnogrids @memb_cons @membl] -Htb #Htb
 * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc -Htb #Htc
 * #td * whd in ⊢ (%→?); #Htd lapply (Htd … Htc) -Htd -Htc #Htd
 * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd #Hte
-whd in ⊢ (%→?); #Htf cases (Htf … Hte) -Htf -Hte 
+* #tf * whd in ⊢ (%→?); #Htf lapply (Htf … Hte) -Htf -Hte #Htf
+whd in ⊢ (%→?); #Htg cases (Htg … Htf) -Htg -Htf
   [* whd in ⊢ ((??%?)→?); #Habs destruct (Habs)]
-* #_ #Htf lapply (Htf (reverse ? l) 〈c,true〉 ls (refl …) (refl …) ?) 
-  [#x #membl @Hnomarks @daemon] -Htf #Htf >Htf >reverse_reverse %
+* #_ #Htg lapply (Htg (〈grid,false〉::reverse ? l) 〈c,true〉 ls (refl …) (refl …) ?) 
+  [#x #membl @Hnomarks @daemon] -Htg #Htg >Htg >reverse_cons >reverse_reverse
+   >associative_append %
 qed.
 
-
 (* init_copy 
 
            init_current_on_match; (* no marks in current *)
@@ -105,10 +105,8 @@ qed.
 *)
 
 definition init_copy ≝ 
-  seq ? init_current_on_match
-    (seq ? (move_r ?) 
-      (seq ? (adv_to_mark_r ? (is_marked ?))
-        (adv_mark_r ?))).
+  init_current_on_match · move_r ? · 
+    adv_to_mark_r ? (is_marked ?) · adv_mark_r ?.
 
 definition R_init_copy ≝ λt1,t2.
   ∀l1,l2,c,ls,d,rs. 
@@ -234,12 +232,8 @@ cases HR -HR
   ]
 qed. *)
 
-include "turing/universal/move_tape.ma".
-
-definition exec_move ≝ 
-  seq ? init_copy
-    (seq ? copy
-      (seq ? (move_r …) move_tape)).
+definition exec_action ≝ 
+  init_copy · copy · move_r … · move_tape.
 
 definition map_move ≝ 
   λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ].
@@ -251,7 +245,7 @@ definition map_move ≝
        mv ≠ null → c1 ≠ null
      dal fatto che c1 e mv sono contenuti nella table
  *)
-definition R_exec_move ≝ λt1,t2.
+definition R_exec_action ≝ λ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〉]) → 
@@ -325,7 +319,7 @@ qed.
 
 *)
 
-lemma sem_exec_move : Realize ? exec_move R_exec_move.
+lemma sem_exec_action : Realize ? exec_action R_exec_action.
 #intape
 cases (sem_seq … sem_init_copy
         (sem_seq … sem_copy
@@ -447,7 +441,7 @@ if is_false(current) (* current state is not final *)
     match_tuple;
     if is_marked(current) = false (* match ok *)
       then 
-           exec_move
+           exec_action
            move_r;
       else sink;
    else nop;
@@ -455,28 +449,26 @@ if is_false(current) (* current state is not final *)
 
 definition uni_step ≝ 
   ifTM ? (test_char STape (λc.\fst c == bit false))
-    (single_finalTM ? (seq ? init_match
-      (seq ? match_tuple
+    (single_finalTM ? 
+      (init_match · match_tuple ·
         (ifTM ? (test_char ? (λc.¬is_marked ? c))
-          (seq ? exec_move (move_r …))
-          (nop ?) (* sink *)
-          tc_true))))
-    (nop ?)
-    tc_true.
+          (exec_action · move_r …)
+          (nop ?) 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) → 
+  ∀n,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv.
+  0 < |table| → table_TM (S n) table → 
   match_in_table (S n) (〈s0,false〉::curconfig) 〈c0,false〉
-    (〈s1,false〉::newconfig) 〈c1,false〉 〈mv,false〉 (〈t0,false〉::table) → 
+    (〈s1,false〉::newconfig) 〈c1,false〉 〈mv,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) → 
+    (curconfig@〈c0,false〉::〈grid,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) ∧
+    (newconfig@〈c2,false〉::〈grid,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).
    
@@ -485,59 +477,61 @@ definition R_uni_step_false ≝ λt1,t2.
   
 axiom sem_match_tuple : Realize ? match_tuple R_match_tuple.
 
+definition us_acc : states ? uni_step ≝ (inr … (inl … (inr … start_nop))).
+
 lemma sem_uni_step :
-  accRealize ? uni_step (inr … (inl … (inr … start_nop)))
+  accRealize ? uni_step us_acc
      R_uni_step_true R_uni_step_false. 
 @(acc_sem_if_app STape … (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 FSUnialpha c)) *)
-          (sem_seq … sem_exec_move (sem_move_r …))
+          (sem_seq … sem_exec_action (sem_move_r …))
           (sem_nop …))))
    (sem_nop …)
    …)
 [@sem_test_char||]
 [ #intape #outtape 
   #ta whd in ⊢ (%→?); #Hta #HR
-  #n #t0 #table #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
+  #n #fulltable #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv
+  #Htable_len cut (∃t0,table. fulltable =〈bar,false〉::〈t0,false〉::table) [(* 0 < |table| *) @daemon]
+  * #t0 * #table #Hfulltable >Hfulltable -fulltable 
   #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 
+  #tb * whd in ⊢ (%→?); #Htb
   lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) s0 t0 ???)
   [ >Hta >associative_append %
-  | (* utilizzare Hmatch
-         cases (match_in_table_to_tuple … Hmatch Htable)
-       ma serve l'iniettività di mk_tuple...
-     *) @daemon
-  | (* idem *) @daemon
+  | @daemon
+  | @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 (con lemma) *) @daemon
+    | (* Hmatch *) @daemon
     | (* Htable *) @daemon
     | (* Htable, Hmatch → |config| = n 
        necessaria modifica in R_match_tuple, le dimensioni non corrispondono
-      *) @daemon ]
+      *) @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
+      >Hnewc #Htd cut (mv1 = 〈mv,false〉)
+      [@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))
+              ((table1@〈bar,false〉::〈s0,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉::
+               newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs))
       [ >Htd @eq_f3 //
         [ >reverse_append >reverse_single %
         | >associative_append >associative_append normalize
-          >associative_append >associative_append <Hmv1 %
+          >associative_append >associative_append >Hmv1 % 
         ]
       ]
       -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte
@@ -547,7 +541,7 @@ lemma sem_uni_step :
       | (* only_bits (〈s1,false〉::newconfig) *) @daemon
       | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon
       | (* no_marks (reverse ? curconfig) *) @daemon
-      | <Hmv1 >Hnewc in Htableeq; 
+      | >Hmv1 in Htableeq; >Hnewc 
         >associative_append >associative_append normalize
         >associative_append >associative_append
         #Htableeq <Htableeq // ]
@@ -556,13 +550,13 @@ lemma sem_uni_step :
       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 ⊢ (???%);
+        change with ((〈bar,false〉::〈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
+        >associative_append normalize >associative_append
+        >Hmv1 % 
+      | @Hliftte
       ]
      | //
      ]
@@ -577,4 +571,4 @@ lemma sem_uni_step :
   #b #Hb cases (Ht1 ? Hb) #Hb' #Ht3 >Ht2 % //
   cases b in Hb'; normalize #H1 //
 ]
-qed.
+qed.
\ No newline at end of file