X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Funi_step.ma;h=b28efc472c44393948b86b78f44a6f1fca352f0b;hb=0460fd3dc2909efe0baa6592281d0cf0527165ff;hp=128c81aa57da708ea32bf4b5298954ce1e799daf;hpb=690675dde36407d039e9d05047bc7909202170c1;p=helm.git diff --git a/matita/matita/lib/turing/universal/uni_step.ma b/matita/matita/lib/turing/universal/uni_step.ma index 128c81aa5..b28efc472 100644 --- a/matita/matita/lib/turing/universal/uni_step.ma +++ b/matita/matita/lib/turing/universal/uni_step.ma @@ -15,6 +15,7 @@ *) include "turing/universal/copy.ma". +include "turing/universal/move_tape.ma". (* @@ -58,23 +59,25 @@ 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 ?)))))). - + (seq ? (move_r ?) + (seq ? (mark ?) + (seq ? (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,15 +86,17 @@ 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. @@ -234,8 +239,6 @@ cases HR -HR ] qed. *) -include "turing/universal/move_tape.ma". - definition exec_move ≝ seq ? init_copy (seq ? copy @@ -465,18 +468,18 @@ definition uni_step ≝ 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) → + ∀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〉 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). @@ -486,52 +489,58 @@ definition R_uni_step_false ≝ λt1,t2. axiom sem_match_tuple : Realize ? match_tuple R_match_tuple. lemma sem_uni_step : - accRealize ? uni_step (inr … (inl … (inr … 0))) + accRealize ? uni_step (inr … (inl … (inr … start_nop))) 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)) +@(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_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 % - | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @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 *) @daemon ] + | (* Htable, Hmatch → |config| = n + necessaria modifica in R_match_tuple, le dimensioni non corrispondono + *) @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 associative_append >associative_append >Hmv1 % ] ] -Htd #Htd lapply (Hte … (S n) … Htd … Ht1') -Htd -Hte @@ -541,7 +550,7 @@ lemma sem_uni_step : | (* only_bits (〈s1,false〉::newconfig) *) @daemon | (* only_bits (curconfig@[〈s0,false〉]) *) @daemon | (* no_marks (reverse ? curconfig) *) @daemon - | Hnewc in Htableeq; + | >Hmv1 in Htableeq; >Hnewc >associative_append >associative_append normalize >associative_append >associative_append #Htableeq 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 associative_append normalize >associative_append % - | >(?: mv = \fst mv1) [| (*Hmatch, Htableeq*) @daemon ] - @Hliftte + >associative_append normalize >associative_append + >Hmv1 % + | @Hliftte ] | // ] @@ -571,4 +580,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