X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fmulti_universal%2Funistep_aux.ma;h=9e240013b01c69b6399f56c79c4f3434e074e158;hb=00101d30f0bf914a26ed2ec61a1aa38598b05209;hp=fa972bdf2b3fb2c7e9daa585f35fa7606c5319c3;hpb=39aab7babf51252cecb81a66af82fe797e8dcbe7;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 fa972bdf2..9e240013b 100644 --- a/matita/matita/lib/turing/multi_universal/unistep_aux.ma +++ b/matita/matita/lib/turing/multi_universal/unistep_aux.ma @@ -13,6 +13,7 @@ include "turing/multi_universal/moves_2.ma". include "turing/multi_universal/match.ma". include "turing/multi_universal/copy.ma". include "turing/multi_universal/alphabet.ma". +include "turing/multi_universal/tuples.ma". (* @@ -51,34 +52,31 @@ include "turing/multi_universal/alphabet.ma". cfg_to_obj *) -definition obj ≝ 0. -definition cfg ≝ 1. -definition prg ≝ 2. +definition obj ≝ (0:DeqNat). +definition cfg ≝ (1:DeqNat). +definition prg ≝ (2:DeqNat). definition obj_to_cfg ≝ - mmove cfg FSUnialpha 2 L · mmove cfg FSUnialpha 2 L · (ifTM ?? (inject_TM ? (test_null ?) 2 obj) - (inject_TM ? (write FSUnialpha (bit false)) 2 cfg · - inject_TM ? (move_r FSUnialpha) 2 cfg · - inject_TM ? (write FSUnialpha (bit false)) 2 cfg) - (inject_TM ? (write FSUnialpha (bit true)) 2 cfg · - inject_TM ? (move_r FSUnialpha) 2 cfg · - copy_step obj cfg FSUnialpha 2) tc_true) · - inject_TM ? (move_l FSUnialpha) 2 cfg · + (copy_step obj cfg FSUnialpha 2 · + mmove cfg FSUnialpha 2 L · + mmove obj FSUnialpha 2 L) + (inject_TM ? (write FSUnialpha null) 2 cfg) + tc_true) · inject_TM ? (move_to_end FSUnialpha L) 2 cfg · - inject_TM ? (move_r FSUnialpha) 2 cfg. + mmove cfg FSUnialpha 2 R. definition R_obj_to_cfg ≝ λt1,t2:Vector (tape FSUnialpha) 3. - ∀c,opt,ls. - nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::opt::ls) (None ?) [ ] → + ∀c,ls. + nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls) (None ?) [ ] → (∀lso,x,rso.nth obj ? t1 (niltape ?) = midtape FSUnialpha lso x rso → t2 = change_vec ?? t1 - (mk_tape ? [ ] (option_hd ? (reverse ? (c::opt::ls))) (tail ? (reverse ? (c::opt::ls)))) cfg) ∧ + (mk_tape ? [ ] (option_hd ? (reverse ? (x::ls))) (tail ? (reverse ? (x::ls)))) cfg) ∧ (current ? (nth obj ? t1 (niltape ?)) = None ? → t2 = change_vec ?? t1 - (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (bit false::bit false::ls))) - (tail ? (reverse ? (bit false :: bit false::ls)))) cfg). + (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (null::ls))) + (tail ? (reverse ? (null::ls)))) cfg). axiom sem_move_to_end_l : ∀sig. move_to_end sig L ⊨ R_move_to_end_l sig. axiom accRealize_to_Realize : @@ -90,120 +88,395 @@ lemma eq_mk_tape_rightof : #alpha #a #al % qed. +axiom daemon : ∀P:Prop.P. + +definition option_cons ≝ λsig.λc:option sig.λl. + match c with [ None ⇒ l | Some c0 ⇒ c0::l ]. + +lemma tape_move_mk_tape_R : + ∀sig,ls,c,rs. + (c = None ? → ls = [ ] ∨ rs = [ ]) → + tape_move ? (mk_tape sig ls c rs) R = + mk_tape ? (option_cons ? c ls) (option_hd ? rs) (tail ? rs). +#sig * [ * [ * | #c * ] | #l0 #ls0 * [ * +[| #r0 #rs0 #H @False_ind cases (H (refl ??)) #H1 destruct (H1) ] | #c * ] ] +normalize // +qed. + lemma sem_obj_to_cfg : obj_to_cfg ⊨ R_obj_to_cfg. @(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?) - (sem_seq ?????? (sem_move_multi ? 2 cfg L ?) (sem_seq ?????? (sem_if ?????????? (sem_test_null_multi ?? obj ?) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_write FSUnialpha (bit false))) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_r ?)) - (sem_inject ???? cfg ? (sem_write FSUnialpha (bit false))))) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_write FSUnialpha (bit true))) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_r ?)) (accRealize_to_Realize … (sem_copy_step …))))) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_l ?)) - (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?)) - (sem_inject ???? cfg ? (sem_move_r ?))))))) // + (sem_seq ?????? (accRealize_to_Realize … (sem_copy_step …)) + (sem_seq ?????? (sem_move_multi ? 2 cfg L ?) + (sem_move_multi ? 2 obj L ?))) + (sem_inject ???? cfg ? (sem_write FSUnialpha null))) + (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?)) + (sem_move_multi ? 2 cfg R ?)))) // #ta #tb * #tc * whd in ⊢ (%→?); #Htc * -#td * whd in ⊢ (%→?); #Htd * -#te * * -[ * #tf * * #Hcurtd #Htf * - #tg * * whd in ⊢ (%→?); #Htg1 #Htg2 * - #th * * * whd in ⊢ (%→%→?); #Hth1 #Hth2 #Hth3 * whd in ⊢ (%→?); - #Hte1 #Hte2 * - #tj * * * #Htj1 #Htj2 #Htj3 * - #tk * * * #Htk1 #Htk2 #Htk3 * whd in ⊢ (%→?); - #Htb1 #Htb2 #c #opt_mark #ls #Hta1 % - [ #lso #x #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc - >Htc in Htd; >nth_change_vec // >change_vec_change_vec - change with (midtape ????) in match (tape_move ???); #Htd >Htd in Htf; #Htf - destruct (Htf) - - - - -lemma wsem_copy : ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → - copy src dst sig n ⊫ R_copy src dst sig n. -#src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop -lapply (sem_while … (sem_copy_step src dst sig n Hneq Hsrc Hdst) … Hloop) // --Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar -[ whd in ⊢ (%→?); * #Hnone #Hout % - [#_ @Hout - |#ls #x #x0 #rs #ls0 #rs0 #Hsrc1 #Hdst1 @False_ind cases Hnone - [>Hsrc1 normalize #H destruct (H) | >Hdst1 normalize #H destruct (H)] - ] -|#tc #td * #x * #y * * #Hcx #Hcy #Htd #Hstar #IH #He lapply (IH He) -IH * - #IH1 #IH2 % - [* [>Hcx #H destruct (H) | >Hcy #H destruct (H)] - |#ls #x' #y' #rs #ls0 #rs0 #Hnth_src #Hnth_dst - >Hnth_src in Hcx; whd in ⊢ (??%?→?); #H destruct (H) - >Hnth_dst in Hcy; whd in ⊢ (??%?→?); #H destruct (H) - >Hnth_src in Htd; >Hnth_dst -Hnth_src -Hnth_dst - cases rs - [(* the source tape is empty after the move *) - #Htd lapply (IH1 ?) - [%1 >Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] >nth_change_vec //] - #Hout (* whd in match (tape_move ???); *) %1 %{([])} %{rs0} % - [% [// | // ] - |whd in match (reverse ??); whd in match (reverse ??); - >Hout >Htd @eq_f2 // cases rs0 // - ] - |#c1 #tl1 cases rs0 - [(* the dst tape is empty after the move *) - #Htd lapply (IH1 ?) [%2 >Htd >nth_change_vec //] - #Hout (* whd in match (tape_move ???); *) %2 %{[ ]} %{(c1::tl1)} % - [% [// | // ] - |whd in match (reverse ??); whd in match (reverse ??); - >Hout >Htd @eq_f2 // - ] - |#c2 #tl2 whd in match (tape_move_mono ???); whd in match (tape_move_mono ???); - #Htd - cut (nth src (tape sig) td (niltape sig)=midtape sig (x::ls) c1 tl1) - [>Htd >nth_change_vec_neq [2:@(not_to_not … Hneq) //] @nth_change_vec //] - #Hsrc_td - cut (nth dst (tape sig) td (niltape sig)=midtape sig (x::ls0) c2 tl2) - [>Htd @nth_change_vec //] - #Hdst_td cases (IH2 … Hsrc_td Hdst_td) -Hsrc_td -Hdst_td - [* #rs01 * #rs02 * * #H1 #H2 #H3 %1 - %{(c2::rs01)} %{rs02} % [% [@eq_f //|normalize @eq_f @H2]] - >Htd in H3; >change_vec_commute // >change_vec_change_vec - >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec - #H >reverse_cons >associative_append >associative_append @H - |* #rs11 * #rs12 * * #H1 #H2 #H3 %2 - %{(c1::rs11)} %{rs12} % [% [@eq_f //|normalize @eq_f @H2]] - >Htd in H3; >change_vec_commute // >change_vec_change_vec - >change_vec_commute [2:@(not_to_not … Hneq) //] >change_vec_change_vec - #H >reverse_cons >associative_append >associative_append @H - ] +#td * * +[ * #te * * #Hcurtc #Hte + * destruct (Hte) #te * * + [ whd in ⊢ (%→%→?); * #x * #y * * -Hcurtc #Hcurtc1 #Hcurtc2 #Hte + * #tf * whd in ⊢ (%→%→?); #Htf #Htd + * #tg * * * whd in ⊢ (%→%→%→%→?); #Htg1 #Htg2 #Htg3 #Htb + #c #ls #Hta1 % + [ #lso #x0 #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof + whd in match (tape_move ???); #Htc + cut (tg = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg) + [@daemon] -Htg1 -Htg2 -Htg3 #Htg destruct (Htg Htf Hte Htd Htc Htb) + >change_vec_change_vec >change_vec_change_vec + >change_vec_commute // >change_vec_change_vec + >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec + >change_vec_commute // >change_vec_change_vec + >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //] + >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //] + >change_vec_commute [|@sym_not_eq //] @eq_f3 // + [ >Hta2 cases rso in Hta2; whd in match (tape_move_mono ???); + [ #Hta2 whd in match (tape_move ???); tape_move_mk_tape_R [| #_ % %] >reverse_cons + >nth_change_vec_neq in Hcurtc1; [|@sym_not_eq //] >Hta2 + normalize in ⊢ (%→?); #H destruct (H) % ] + | #Hta2 >Htc in Hcurtc1; >nth_change_vec_neq [| @sym_not_eq //] + >Hta2 #H destruct (H) ] + | * #Hcurtc0 #Hte #_ #_ #c #ls #Hta1 >Hta1 in Htc; >eq_mk_tape_rightof + whd in match (tape_move ???); #Htc >Htc in Hcurtc0; * + [ >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //] + #Hcurtc #Hcurtc0 >Hcurtc0 in Hcurtc; * #H @False_ind @H % + | >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H) ] ] +| * #te * * #Hcurtc #Hte + * whd in ⊢ (%→%→?); #Htd1 #Htd2 + * #tf * * * #Htf1 #Htf2 #Htf3 whd in ⊢ (%→?); #Htb + #c #ls #Hta1 % + [ #lso #x #rso #Hta2 >Htc in Hcurtc; >nth_change_vec_neq [|@sym_not_eq //] + >Hta2 normalize in ⊢ (%→?); #H destruct (H) + | #_ >Hta1 in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc + destruct (Hte) cut (td = change_vec ?? tc (midtape ? ls null []) cfg) + [@daemon] -Htd1 -Htd2 #Htd + -Htf1 cut (tf = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[null])) cfg) + [@daemon] -Htf2 -Htf3 #Htf destruct (Htf Htd Htc Htb) + >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec + >change_vec_change_vec >change_vec_change_vec >nth_change_vec // + >reverse_cons >tape_move_mk_tape_R /2/ ] +] +qed. + +definition test_null_char ≝ test_char FSUnialpha (λc.c == null). + +definition R_test_null_char_true ≝ λt1,t2. + current FSUnialpha t1 = Some ? null ∧ t1 = t2. + +definition R_test_null_char_false ≝ λt1,t2. + current FSUnialpha t1 ≠ Some ? null ∧ t1 = t2. + +lemma sem_test_null_char : + test_null_char ⊨ [ tc_true : R_test_null_char_true, R_test_null_char_false]. +#t1 cases (sem_test_char FSUnialpha (λc.c == null) t1) #k * #outc * * #Hloop #Htrue +#Hfalse %{k} %{outc} % [ % +[ @Hloop +| #Houtc cases (Htrue ?) [| @Houtc] * #c * #Hcurt1 #Hcnull lapply (\P Hcnull) + -Hcnull #H destruct (H) #Houtc1 % + [ @Hcurt1 | Hcurt1 in Hc; #Hc lapply (Hc ? (refl ??)) + >(?:((null:FSUnialpha) == null) = true) [|@(\b (refl ??)) ] + #H destruct (H) + | Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc + cut (te = tc) [@daemon] -Hte1 -Hte2 #Hte + cut (tf = change_vec ? 3 te (mk_tape ? [ ] (None ?) (reverse ? ls@[c])) cfg) + [@daemon] -Htf1 -Htf2 -Htf3 #Htf + destruct (Htf Hte Htc Htb) + >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec + >nth_change_vec // >tape_move_mk_tape_R [| #_ % % ] + >reverse_cons % + | #Hc >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc + >Htc in Hcurtc; >nth_change_vec // normalize in ⊢ (%→?); + #H destruct (H) @False_ind cases Hc /2/ ] +| * #te * * * #Hcurtc #Hte1 #Hte2 + * #tf * * + [ (* purtroppo copy_step assume che la destinazione sia Some (almeno come semantica) *) + STOP + * #x * #y * * #Hcurte_cfg #Hcurte_obj #Htf + * #tg * whd in ⊢ (%→%→?); #Htg #Htd + * #th * * * #Hth1 #Hth2 #Hth3 + whd in ⊢ (%→%); #Htb + #c #ls #Hta % #Hc + [ >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc + >Htc in Hcurtc; >nth_change_vec // normalize in ⊢ (%→?); >Hc + * #H @False_ind /2/ + | >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc + cut (te = tc) [@daemon] -Hte1 -Hte2 #Hte +qed. +*) + +(* macchina che muove il nastro obj a destra o sinistra a seconda del valore + del current di prg, che codifica la direzione in cui ci muoviamo *) + +definition char_to_move ≝ λc.match c with + [ bit b ⇒ if b then R else L + | _ ⇒ N]. + +definition char_to_bit_option ≝ λc.match c with + [ bit b ⇒ Some ? (bit b) + | _ ⇒ None ?]. -lemma terminate_copy : ∀src,dst,sig,n,t. - src ≠ dst → src < S n → dst < S n → copy src dst sig n ↓ t. -#src #dst #sig #n #t #Hneq #Hsrc #Hdts -@(terminate_while … (sem_copy_step …)) // -<(change_vec_same … t src (niltape ?)) -cases (nth src (tape sig) t (niltape ?)) -[ % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct -|2,3: #a0 #al0 % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); #Hx destruct -| #ls #c #rs lapply c -c lapply ls -ls lapply t -t elim rs - [#t #ls #c % #t1 * #x * #y * * >nth_change_vec // normalize in ⊢ (%→?); - #H1 destruct (H1) #_ >change_vec_change_vec #Ht1 % - #t2 * #x0 * #y0 * * >Ht1 >nth_change_vec_neq [|@sym_not_eq //] - >nth_change_vec // normalize in ⊢ (%→?); #H destruct (H) - |#r0 #rs0 #IH #t #ls #c % #t1 * #x * #y * * >nth_change_vec // - normalize in ⊢ (%→?); #H destruct (H) #Hcur - >change_vec_change_vec >change_vec_commute // #Ht1 >Ht1 @IH - ] -] +definition tape_move_obj : mTM FSUnialpha 2 ≝ + ifTM ?? + (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit false)) 2 prg) + (mmove obj FSUnialpha 2 L) + (ifTM ?? + (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit true)) 2 prg) + (mmove obj FSUnialpha 2 R) + (nop ??) + tc_true) + tc_true. + +definition restart_tape ≝ λi. + inject_TM ? (move_to_end FSUnialpha L) 2 i · + mmove i FSUnialpha 2 R. + +definition unistep ≝ + 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 legal_tape ≝ λn,l,h,t. + ∃state,char,table. + nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) → + is_config n (bar::state@[char]) → + nth prg ? t1 (niltape ?) = midtape ? [ ] bar table → + bar::table = table_TM n l h → *) + +definition list_of_tape ≝ λsig,t. + left sig t@option_cons ? (current ? t) (right ? t). + +definition low_char' ≝ λc. + match c with + [ None ⇒ null + | Some b ⇒ if (is_bit b) then b else null + ]. + +lemma low_char_option : ∀s. + low_char' (option_map FinBool FSUnialpha bit s) = low_char s. +* // qed. -lemma sem_copy : ∀src,dst,sig,n. - src ≠ dst → src < S n → dst < S n → - copy src dst sig n ⊨ R_copy src dst sig n. -#i #j #sig #n #Hneq #Hi #Hj @WRealize_to_Realize [/2/| @wsem_copy // ] +definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3. + ∀state,char,table. + (* cfg *) + nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) → + is_config n (bar::state@[char]) → + (* prg *) + nth prg ? t1 (niltape ?) = midtape ? [ ] bar table → + bar::table = table_TM n l h → + (* obj *) + only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) → + let conf ≝ (bar::state@[char]) in + (∃ll,lr.bar::table = ll@conf@lr) → +(* + ∃nstate,nchar,m,t. tuple_encoding n h t = (conf@nstate@[nchar;m]) ∧ + mem ? t l ∧ *) + ∀nstate,nchar,m,t. + tuple_encoding n h t = (conf@nstate@[nchar;m])→ + mem ? t l → + let new_obj ≝ + tape_move_mono ? (nth obj ? t1 (niltape ?)) + 〈char_to_bit_option nchar,char_to_move m〉 in + let next_char ≝ low_char' (current ? new_obj) in + t2 = + change_vec ?? + (change_vec ?? t1 (midtape ? [ ] bar (nstate@[next_char])) cfg) + new_obj obj. + +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)). + +lemma map_list_of_tape: ∀A,B,f,t. + list_of_tape B (tape_map ?? f t) = map ?? f (list_of_tape A t). +#A #B #f * // normalize // #ls #c #rs map_write >map_move % +qed. + +lemma map_move_mono: ∀t,cout,m. + tape_move_mono ? (tape_map FinBool ? bit t) + 〈char_to_bit_option (low_char cout), char_to_move (low_mv m)〉 + = tape_map ?? bit (tape_move_mono ? t 〈cout,m〉). +@map_action +qed. + +definition R_unistep_high ≝ λM:normalTM.λc:nconfig (no_states M).λt1,t2. + t1 = low_tapes M c → + t2 = low_tapes M (step ? M c). + +lemma R_unistep_equiv : ∀M,c,t1,t2. + R_unistep (no_states M) (graph_enum ?? (ntrans M)) (nhalt M) t1 t2 → + R_unistep_high M c t1 t2. +#M #c #t1 #t2 #H #Ht1 +lapply (initial_bar ? (nhalt M) (graph_enum ?? (ntrans M)) (nTM_nog ?)) #Htable +(* tup = current tuple *) +cut (∃t.t = 〈〈cstate … c,current ? (ctape … c)〉, + ntrans M 〈cstate … c,current ? (ctape … c)〉〉) [% //] * #tup #Htup +(* tup is in the graph *) +cut (mem ? tup (graph_enum ?? (ntrans M))) + [@memb_to_mem >Htup @(graph_enum_complete … (ntrans M)) %] #Hingraph +(* tupe target = 〈qout,cout,m〉 *) +lapply (decomp_target ? (ntrans M 〈cstate … c,current ? (ctape … c)〉)) +* #qout * #cout * #m #Htg >Htg in Htup; #Htup +(* new config *) +cut (step FinBool M c = mk_config ?? qout (tape_move ? (tape_write ? (ctape … c) cout) m)) + [>(config_expand … c) whd in ⊢ (??%?); (* >Htg ?? why not?? *) + cut (trans ? M 〈cstate … c, current ? (ctape … c)〉 = 〈qout,cout,m〉) [Heq1 %] #Hstep +(* new state *) +cut (cstate ?? (step FinBool M c) = qout) [>Hstep %] #Hnew_state +(* new tape *) +cut (ctape ?? (step FinBool M c) = tape_move ? (tape_write ? (ctape … c) cout) m) + [>Hstep %] #Hnew_tape +lapply(H (bits_of_state ? (nhalt M) (cstate ?? c)) + (low_char (current ? (ctape ?? c))) + (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M))) + ??????) +[Htable1 @eq_f Htup + whd in ⊢ (??%?); @eq_f >associative_append % +|>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 +|>Ht1 % +|%{(bits_of_state ? (nhalt M) (cstate ?? c))} %{(low_char (current ? (ctape ?? c)))} + % [% [% [// | cases (current ??) normalize [|#b] % #Hd destruct (Hd)] + |>length_map whd in match (length ??); @eq_f //] + |//] +|>Ht1 >cfg_low_tapes //] -H #H +lapply(H (bits_of_state … (nhalt M) qout) (low_char … cout) + (low_mv … m) tup ? Hingraph) + [>Htup whd in ⊢ (??%?); @eq_f >associative_append %] -H +#Ht2 >Ht2 @(eq_vec ? 3 … (niltape ?)) #i #Hi +cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi + [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi + [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi + [@False_ind /2/ + |>Hi >obj_low_tapes >nth_change_vec // + >Ht1 >obj_low_tapes >Hstep @map_action + ] + |>Hi >cfg_low_tapes >nth_change_vec_neq + [|% whd in ⊢ (??%?→?); #H destruct (H)] + >nth_change_vec // >Hnew_state @eq_f @eq_f >Hnew_tape + @eq_f2 [|2:%] >Ht1 >obj_low_tapes >map_move_mono >low_char_current % + ] + |(* program tapes do not change *) + >Hi >prg_low_tapes + >nth_change_vec_neq [|% whd in ⊢ (??%?→?); #H destruct (H)] + >nth_change_vec_neq [|% whd in ⊢ (??%?→?); #H destruct (H)] + >Ht1 >prg_low_tapes // + ] +qed. + + + + \ No newline at end of file