X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fmulti_universal%2Funistep_aux.ma;h=8033c222627b6123481c324d528aab44652260aa;hb=685c36442ffed93a7bb0de464d35478821884c77;hp=701e59fdc02760e6511a542606dca0a6d71e8f9e;hpb=3447747453bbf439d071d21dcb68149cae3a9068;p=helm.git diff --git a/matita/matita/lib/turing/multi_universal/unistep_aux.ma b/matita/matita/lib/turing/multi_universal/unistep_aux.ma index 701e59fdc..8033c2226 100644 --- a/matita/matita/lib/turing/multi_universal/unistep_aux.ma +++ b/matita/matita/lib/turing/multi_universal/unistep_aux.ma @@ -9,9 +9,8 @@ \ / GNU General Public License Version 2 V_____________________________________________________________*) -include "turing/multi_universal/moves_2.ma". -include "turing/multi_universal/match.ma". -include "turing/multi_universal/copy.ma". +include "turing/auxiliary_machines.ma". +include "turing/auxiliary_multi_machines.ma". include "turing/multi_universal/alphabet.ma". include "turing/multi_universal/tuples.ma". @@ -39,9 +38,11 @@ definition R_obj_to_cfg ≝ λt1,t2:Vector (tape FSUnialpha) 3. (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (null::ls))) (tail ? (reverse ? (null::ls)))) cfg). +(* 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. @@ -58,16 +59,6 @@ lemma tape_move_mk_tape_R : normalize // qed. -lemma eq_vec_change_vec : ∀sig,n.∀v1,v2:Vector sig n.∀i,t,d. - nth i ? v2 d = t → - (∀j.i ≠ j → nth j ? v1 d = nth j ? v2 d) → - v2 = change_vec ?? v1 t i. -#sig #n #v1 #v2 #i #t #d #H1 #H2 @(eq_vec … d) -#i0 #Hlt cases (decidable_eq_nat i0 i) #Hii0 -[ >Hii0 >nth_change_vec // -| >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @H2 @sym_not_eq // ] -qed. - lemma None_or_Some: ∀A.∀a. a =None A ∨ ∃b. a = Some ? b. #A * /2/ #a %2 %{a} % qed. @@ -466,18 +457,29 @@ lemma sem_restart_tape : ∀i,n.i < S n → restart_tape i n ⊨ R_restart_tape * #td * * * #Htd1 #Htd2 #Htd3 whd in ⊢ (%→?); #Htb * [ #Hta_i Htc >nth_change_vec // + | @Htd3 ] ] (* >Htc in Htd1; >nth_change_vec // *) -Htd1 -Htd2 -Htd3 #Htd >Htd in Htb; >Htc >change_vec_change_vec >nth_change_vec // #Htb >Htb % | #r0 #rs0 #Hta_i Htc >nth_change_vec // + | @Htd3 ] ] (* >Htc in Htd1; >nth_change_vec // *) -Htd1 -Htd2 -Htd3 #Htd >Htd in Htb; >Htc >change_vec_change_vec >nth_change_vec // #Htb >Htb % | #l0 #ls0 #Hta_i Htc >nth_change_vec // + | #j #Hij >nth_change_vec_neq // @Htd3 // ]] #Htd >Htd in Htb; >Htc >change_vec_change_vec >change_vec_change_vec >nth_change_vec // #Htb >Htb <(reverse_reverse ? ls0) in ⊢ (???%); cases (reverse ? ls0) @@ -490,13 +492,19 @@ whd in ⊢ (%→?); #Htb * whd in ⊢ (??%?); @eq_f >reverse_reverse normalize >append_nil % ] % ] | * [ #c #rs #Hta_i Htc >nth_change_vec // + | @Htd3 ] ] (* >Htc in Htd1; >nth_change_vec // *) -Htd1 -Htd2 -Htd3 #Htd >Htd in Htb; >Htc >change_vec_change_vec >nth_change_vec // #Htb >Htb % | #l0 #ls0 #c #rs #Hta_i Htc >nth_change_vec // + | @Htd3 ] ] #Htd >Htd in Htb; >Htc >change_vec_change_vec >change_vec_change_vec >nth_change_vec // #Htb >Htb <(reverse_reverse ? ls0) in ⊢ (???%); cases (reverse ? ls0)