X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fmulti_universal%2Funistep.ma;h=ee594814189e3f790076129f75a4d39b4dad9556;hb=d6b8021e8c83eb19033cad0aeaeebf95b327e78a;hp=2e4e1e253d2425dd7a237b6930f1f03628d8d773;hpb=3026dfea8eae158433ed13df5156a733fa926794;p=helm.git diff --git a/matita/matita/lib/turing/multi_universal/unistep.ma b/matita/matita/lib/turing/multi_universal/unistep.ma index 2e4e1e253..ee5948141 100644 --- a/matita/matita/lib/turing/multi_universal/unistep.ma +++ b/matita/matita/lib/turing/multi_universal/unistep.ma @@ -11,6 +11,7 @@ include "turing/multi_universal/unistep_aux.ma". +include "turing/multi_universal/match.ma". definition exec_move ≝ cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg. @@ -101,6 +102,13 @@ definition legal_tape ≝ λn,l,h,t. is_config n (bar::state@[char]) → nth prg ? t1 (niltape ?) = midtape ? [ ] bar table → bar::table = table_TM n l h → *) + +definition deterministic_tuples ≝ λn,h,l. + ∀t1,t2,conf,out1,out2. + is_config n conf → + mem ? t1 l → mem ? t2 l → + tuple_encoding n h t1 = conf@out1 → + tuple_encoding n h t2 = conf@out2 → out1 = out2. definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3. ∀state,char,table. @@ -112,6 +120,8 @@ definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3. bar::table = table_TM n l h → (* obj *) only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) → + (* deterministic tuples *) + deterministic_tuples n h l → let conf ≝ (bar::state@[char]) in (∃ll,lr.bar::table = ll@conf@lr) → (* @@ -148,12 +158,29 @@ definition R_copy_strict ≝ (tail sig rs2)) src) (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)). -axiom sem_copy_strict : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n → +lemma sem_copy_strict : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n → copy src dst sig n ⊨ R_copy_strict src dst sig n. +#src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_copy …)) // +#ta #tb * #Htb1 #Htb2 % [ @Htb1 ] +#ls #x #x0 #rs #ls0 #rs0 #Htamid_src #Htamid_dst #Hlen +cases (Htb2 … Htamid_src Htamid_dst) -Htb1 -Htb2 +[ * #rs1 * #rs2 * * #Hrs0 #Heq #Htb Hrs0 >length_append + >(plus_n_O (|rs1|)) #Hlen cut (|rs2| ≤ 0) [ /2 by le_plus_to_le/ ] + #Hlenrs2 cut (rs2 = [ ]) + [ cases rs2 in Hlenrs2; // #r3 #rs3 normalize #H @False_ind cases (not_le_Sn_O (|rs3|)) /2/ ] + #Hrs2 destruct (Hrs2) >append_nil in Hrs0; #Hrs0 destruct (Hrs0) -Hlenrs2 -Hlen + append_nil % [ % // ] @Htb +| * #rs1 * #rs2 #H %{rs1} %{rs2} @H ] +qed. -axiom daemon : ∀P:Prop.P. +lemma config_to_len : ∀n,b,q,c.is_config n (b::q@[c]) → |q| = S n. +#n #b #q #c * #q0 * #cin * * * #_ #_ #Hq0 #H >(?:q = q0) // +lapply (cons_injective_r ????? H) #H1 @(append_l1_injective … H1) +lapply (eq_f … (length ?) … H) normalize >length_append >length_append +Hta_cfg #H cases (H ?? (refl ??)) -H (* prg starts with a bar, so it's not empty *) #_ >Hta_prg #H lapply (H ??? (refl ??)) -H * -[| cases Htotaltable #ll * #lr #H >H +[| cases Hmatching #ll * #lr #H >H #Hfalse @False_ind cases (Hfalse ll lr) #H1 @H1 //] -* #ll * #lr * #Hintable -Htotaltable #Htc +* #ll * #lr * #Hintable #Htc * #td * whd in ⊢ (%→?); >Htc >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg) #Htd lapply (Htd ? (refl ??)) -Htd @@ -183,61 +210,79 @@ whd in ⊢ (???(???(????%?)??)→?); whd in match (tail ??); #Htd (* move cfg to R *) * #te * whd in ⊢ (%→?); >Htd >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec ->nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec // +>nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec [2:@leb_true_to_le %] +cut (∃ll1.ll@[bar] = bar::ll1) +[ cases ll in Hintable; [ #_ %{[ ]} % ] + #llhd #lltl normalize #H destruct (H) %{(lltl@[bar])} % ] * #ll1 #Hll1 >Htable in Hintable; #Hintable #Hte (* copy *) -lapply (cfg_in_table_to_tuple ???? Hcfg ?? Hintable) -* #newconfig * #m0 * #lr0 * * #Hlr destruct (Hlr) #Hnewcfg #Hm0 +lapply (table_to_list ???? Hcfg ?? Hintable) * #out * #lr0 * #t0 +* * #Hlr #Htuple_t0 #mem_t0 +cut (out = nstate@[nchar;m]) + [@(Hdeterm … Hcfg mem_t0 Hmatch Htuple_t0 Htuple)] #Hout +>(append_cons ? nchar) in Htuple; #Htuple +lapply (tuple_to_config ?????? Hcfg … Htuple) #newconfig cut (∃fo,so.state = fo::so ∧ |so| = n) -[ @daemon ] * #fo * #so * #Hstate_exp #Hsolen -cut (∃fn,sn,cn.newconfig = fn::sn@[cn] ∧ |sn| = n) -[ @daemon ] * #fn * #sn * #cn * #Hnewstate_exp #Hsnlen +[ lapply (config_to_len … Hcfg) cases state [ normalize #H destruct (H) ] + #fo #so normalize #H destruct (H) %{fo} %{so} % // ] +* #fo * #so * #Hstate_exp #Hsolen +cut (∃fn,sn.nstate = fn::sn ∧ |sn| = n) +[ lapply (config_to_len … newconfig) cases nstate [ normalize #H destruct (H) ] + #fn #sn normalize #H destruct (H) %{fn} %{sn} % // ] +* #fn * #sn * #Hnewstate_exp #Hsnlen +>Hstate_exp >Hout in Hlr; >Hnewstate_exp whd in ⊢ (???%→?); #Hlr * #tf * * #_ >Hte >(nth_change_vec ?????? lt_prg) >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg) ->Hstate_exp >Hnewstate_exp +>Hstate_exp >Hnewstate_exp >Hlr whd in match (mk_tape ????); whd in match (tape_move ???); #Htf cases (Htf ?????? (refl ??) (refl ??) ?) -Htf [| whd in match (tail ??); >length_append >length_append - >Hsolen >length_append >Hsnlen //] + >Hsolen >length_append >Hsnlen normalize //] #rs1 * #rs2 whd in match (tail ??); * * #Hrs1rs2 #Hrs1len >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec >append_nil #Htf (* exec_move *) -cut ((sn@[cn])=rs1) +>append_cons in Hrs1rs2; >associative_append #Hrs1rs2 +cut ((sn@[nchar])=rs1) [@(append_l1_injective ?????? Hrs1rs2) >Hrs1len >length_append >length_append normalize >Hsolen >Hsnlen % ] #Hrs1 -cut (m0::lr0=rs2) [@(append_l2_injective ?????? Hrs1rs2) >Hrs1 % ] #Hrs2 -cut (∃ll1. ll@[bar] = bar::ll1) - [@daemon (* ll is at the begininng of a table *)] * #ll1 #Hll +cut (m::lr0=rs2) [@(append_l2_injective ?????? Hrs1rs2) >Hrs1 % ] #Hrs2 whd in ⊢ (%→?); >Htf >nth_change_vec_neq [2:@eqb_false_to_not_eq %] >nth_change_vec [2:@leb_true_to_le %] >nth_change_vec [2:@leb_true_to_le %] >nth_change_vec_neq [2:@eqb_false_to_not_eq %] >nth_change_vec_neq [2:@eqb_false_to_not_eq %] >append_cons reverse_append >reverse_single -(append_cons … bar ll) >Hll >reverse_cons +(append_cons … bar ll) >Hll1 >reverse_cons #sem_exec_move lapply - (sem_exec_move ? m0 ? - (([cn]@reverse FSUnialpha sn) + (sem_exec_move ? m ? + (([nchar]@reverse FSUnialpha sn) @fn::reverse FSUnialpha (ll1@(fo::so)@[char])) lr0 … (refl ??) ????) - [@Hm0 - |@daemon (* OK *) + [ cut (tuple_TM (S n) (tuple_encoding n h t)) // >Htuple + * #qin * #cin * #qout * #cout * #mv * * * * #_ #Hmv #_ #_ + normalize >(?: bar::qin@cin::qout@[cout;mv] = (bar::qin@cin::qout@[cout])@[mv]) + [| normalize >associative_append normalize >associative_append % ] + >(?: bar::(state@[char])@(nstate@[nchar])@[m] = (bar::(state@[char])@(nstate@[nchar]))@[m]) + [|normalize >associative_append >associative_append >associative_append >associative_append >associative_append % ] + #Heq lapply (append_l2_injective_r ?????? Heq) // #H destruct (H) // + | cases newconfig #qout * #cout * * * #_ #Hcout #_ #H destruct (H) -H + lapply (append_l2_injective_r ?????? e0) // #H destruct (H) @Hcout |@Hbits_obj |whd in ⊢ (??%?); >associative_append >associative_append - >associative_append % + >associative_append % |#Htb >Htb @(eq_vec … (niltape ?)) (* tape by tape *) #i #lei2 cases (le_to_or_lt_eq … (le_S_S_to_le …lei2)) [#lei1 cases (le_to_or_lt_eq … (le_S_S_to_le …lei1)) [#lei0 lapply(le_n_O_to_eq … (le_S_S_to_le …lei0)) #eqi nth_change_vec [2:@leb_true_to_le %] - (* dimostrare che la tabella e' univoca - da cui m = m0 e nchar = cn *) + >nth_change_vec [2:@leb_true_to_le %] % |(* cfg tape *) #eqi >eqi >nth_change_vec_neq [2:@eqb_false_to_not_eq %] - >nth_change_vec [2:@leb_true_to_le %] + >nth_change_vec [2:@leb_true_to_le %] + whd in ⊢ (??%?); cut (∀A.∀l:list A.[]@l = l) [//] #Hnil >Hnil + >reverse_append >reverse_single >reverse_reverse % (* idem *) ] |(* prg tape *) #eqi >eqi @@ -247,61 +292,33 @@ lapply >Htable >Hintable >reverse_append >reverse_cons >reverse_reverse >reverse_cons >reverse_reverse >associative_append >associative_append >associative_append - >(append_cons ? bar ll) >Hll @eq_f @eq_f Hnewstate_exp % + >(append_cons ? bar ll) >Hll1 @eq_f @eq_f Hnewstate_exp >Hlr normalize >associative_append >associative_append % ] ] - - qed. - - - cut ((mk_tape FSUnialpha [] - (option_hd FSUnialpha - (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar]))) - (tail FSUnialpha - (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar])))) = - midtape ? [ ] bar (fn::sn@[cn;m0])) - [cut (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar]) = - bar::fn::sn@[cn;m0]) - [>reverse_cons whd in ⊢ (??(??(??%)?)?); >reverse_append >reverse_reverse - >append_cons in ⊢ (???%); % ] #Hrev >Hrev % ] #Hmk_tape >Hmk_tape -Hmk_tape - >change_vec_commute - - >reverse_append >reverse_append - check reverse_cons - reverse_cons - -Htg #Htg - ->(change_vec_commute ????? cfg prg) [2:@eqb_false_to_not_eq %] ->nth_change_vec_neq [2:@eqb_false_to_not_eq %] ->nth_change_vec_neq [2:@eqb_false_to_not_eq %] -lapply (append_l1_injective ?????? Hrs1rs2) -[ >Hsnlen >Hrs1len >length_append >length_append >length_append >length_append - normalize >Hsolen >Hsnlen % ] -#Hrs1 reverse_append #Htg cases (Htg ?? (refl ??)) -Htg -cases m0 - [#mv #_ #Htg #_ - - - - - -[ * - - match_m cfg prg FSUnialpha 2 · - restart_tape cfg · copy prg cfg FSUnialpha 2 · - cfg_to_obj · tape_move_obj · restart_tape prg · obj_to_cfg. - definition tape_map ≝ λA,B:FinSet.λf:A→B.λt. mk_tape B (map ?? f (left ? t)) (option_map ?? f (current ? t)) (map ?? f (right ? t)). - + +(* da spostare *) +lemma map_reverse: ∀A,B,f,l. + map ?? f (reverse A l) = reverse B (map ?? f l). +#A #B #f #l elim l // +#a #l1 #Hind >reverse_cons >reverse_cons rev_append_def >rev_append_def >append_nil + >append_nil rev_append_def >rev_append_def + >append_nil >append_nil Htable1 @eq_f Htup whd in ⊢ (??%?); @eq_f >associative_append % +|#tx #ty #conf #outx #outy #isconf #memx #memy #tuplex #tupley + @(deterministic M … (refl ??) memx memy isconf tuplex tupley) |>Ht1 >obj_low_tapes >map_list_of_tape elim (list_of_tape ??) [#b @False_ind | #b #tl #Hind #a * [#Ha >Ha //| @Hind]] |@sym_eq @Htable @@ -432,4 +451,8 @@ cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi >nth_change_vec_neq [|% whd in ⊢ (??%?→?); #H destruct (H)] >Ht1 >prg_low_tapes // ] -qed. +qed. + +lemma sem_unistep : ∀M.unistep ⊨ R_unistep_high M. +#M @(Realize_to_Realize … (R_unistep_equiv …)) // +qed.