X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fmulti_universal%2Fmatch.ma;h=9cd11f61007dce904310557b583ee95ba990e41c;hb=143c97a8fe657eb7b041dec2747b0e67b5899762;hp=513de456217a9566136417c9fb5e4f702b95c243;hpb=dabd7add16b4e678f48bc15cd0d992b80fbc9216;p=helm.git diff --git a/matita/matita/lib/turing/multi_universal/match.ma b/matita/matita/lib/turing/multi_universal/match.ma index 513de4562..9cd11f610 100644 --- a/matita/matita/lib/turing/multi_universal/match.ma +++ b/matita/matita/lib/turing/multi_universal/match.ma @@ -12,107 +12,13 @@ (* *) (**************************************************************************) -include "turing/multi_universal/compare.ma". -include "turing/multi_universal/par_test.ma". -include "turing/multi_universal/moves_2.ma". +include "turing/auxiliary_multi_machines.ma". -definition Rtc_multi_true ≝ - λalpha,test,n,i.λt1,t2:Vector ? (S n). - (∃c. current alpha (nth i ? t1 (niltape ?)) = Some ? c ∧ test c = true) ∧ t2 = t1. - -definition Rtc_multi_false ≝ - λalpha,test,n,i.λt1,t2:Vector ? (S n). - (∀c. current alpha (nth i ? t1 (niltape ?)) = Some ? c → test c = false) ∧ t2 = t1. - -lemma sem_test_char_multi : - ∀alpha,test,n,i.i ≤ n → - inject_TM ? (test_char ? test) n i ⊨ - [ tc_true : Rtc_multi_true alpha test n i, Rtc_multi_false alpha test n i ]. -#alpha #test #n #i #Hin #int -cases (acc_sem_inject … Hin (sem_test_char alpha test) int) -#k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ % -[ @Hloop -| #Hqtrue lapply (Htrue Hqtrue) * * * #c * - #Hcur #Htestc #Hnth_i #Hnth_j % - [ %{c} % // - | @(eq_vec … (niltape ?)) #i0 #Hi0 - cases (decidable_eq_nat i0 i) #Hi0i - [ >Hi0i @Hnth_i - | @sym_eq @Hnth_j @sym_not_eq // ] ] ] -| #Hqfalse lapply (Hfalse Hqfalse) * * #Htestc #Hnth_i #Hnth_j % - [ @Htestc - | @(eq_vec … (niltape ?)) #i0 #Hi0 - cases (decidable_eq_nat i0 i) #Hi0i - [ >Hi0i @Hnth_i - | @sym_eq @Hnth_j @sym_not_eq // ] ] ] -qed. - -definition Rm_test_null_true ≝ - λalpha,n,i.λt1,t2:Vector ? (S n). - current alpha (nth i ? t1 (niltape ?)) ≠ None ? ∧ t2 = t1. - -definition Rm_test_null_false ≝ - λalpha,n,i.λt1,t2:Vector ? (S n). - current alpha (nth i ? t1 (niltape ?)) = None ? ∧ t2 = t1. - -lemma sem_test_null_multi : ∀alpha,n,i.i ≤ n → - inject_TM ? (test_null ?) n i ⊨ - [ tc_true : Rm_test_null_true alpha n i, Rm_test_null_false alpha n i ]. -#alpha #n #i #Hin #int -cases (acc_sem_inject … Hin (sem_test_null alpha) int) -#k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ % -[ @Hloop -| #Hqtrue lapply (Htrue Hqtrue) * * #Hcur #Hnth_i #Hnth_j % // - @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) #Hi0i - [ >Hi0i @sym_eq @Hnth_i | @sym_eq @Hnth_j @sym_not_eq // ] ] -| #Hqfalse lapply (Hfalse Hqfalse) * * #Hcur #Hnth_i #Hnth_j % - [ @Hcur - | @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) // - #Hi0i @sym_eq @Hnth_j @sym_not_eq // ] ] -qed. - -definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n. - match (nth src (option sig) v (None ?)) with - [ None ⇒ false - | Some x ⇒ notb (nth dst (DeqOption sig) v (None ?) == None ?) ]. - -definition mmove_states ≝ initN 2. - -definition mmove0 : mmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)). -definition mmove1 : mmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)). - -definition trans_mmove ≝ - λi,sig,n,D. - λp:mmove_states × (Vector (option sig) (S n)). - let 〈q,a〉 ≝ p in match (pi1 … q) with - [ O ⇒ 〈mmove1,change_vec ? (S n) (null_action ? n) (〈None ?,D〉) i〉 - | S _ ⇒ 〈mmove1,null_action sig n〉 ]. - -definition mmove ≝ - λi,sig,n,D. - mk_mTM sig n mmove_states (trans_mmove i sig n D) - mmove0 (λq.q == mmove1). - -definition Rm_multi ≝ - λalpha,n,i,D.λt1,t2:Vector ? (S n). - t2 = change_vec ? (S n) t1 (tape_move alpha (nth i ? t1 (niltape ?)) D) i. - -lemma sem_move_multi : - ∀alpha,n,i,D.i ≤ n → - mmove i alpha n D ⊨ Rm_multi alpha n i D. -#alpha #n #i #D #Hin #int %{2} -%{(mk_mconfig ? mmove_states n mmove1 ?)} -[| % - [ whd in ⊢ (??%?); @eq_f whd in ⊢ (??%?); @eq_f % - | whd >tape_move_multi_def - <(change_vec_same … (ctapes …) i (niltape ?)) - >pmap_change tape_move_null_action % ] ] - qed. - +(* rewind *) definition rewind ≝ λsrc,dst,sig,n. parmove src dst sig n L · mmove src sig n R · mmove dst sig n R. -definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n). +definition R_rewind_strong ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n). (∀x,x0,xs,rs. nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → ∀ls0,y,y0,target,rs0.|xs| = |target| → @@ -120,32 +26,42 @@ definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n). outt = change_vec ?? (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src) (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧ + (∀x,x0,xs,rs. + nth dst ? int (niltape ?) = midtape sig (xs@[x0]) x rs → + ∀ls0,y,y0,target,rs0.|xs| = |target| → + nth src ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → + outt = change_vec ?? + (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) dst) + (midtape sig ls0 y0 (reverse ? target@y::rs0)) src) ∧ (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → + outt = int) ∧ + (∀x,rs.nth dst ? int (niltape ?) = midtape sig [] x rs → + ∀ls0,y,rs0.nth src ? int (niltape ?) = midtape sig ls0 y rs0 → outt = int). -(* -theorem accRealize_to_Realize : - ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc. - M ⊨ [ acc: Rtrue, Rfalse ] → M ⊨ Rtrue ∪ Rfalse. -#sig #n #M #Rtrue #Rfalse #acc #HR #t -cases (HR t) #k * #outc * * #Hloop -#Htrue #Hfalse %{k} %{outc} % // -cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase -[ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ] -qed. -*) +definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n). + (∀x,x0,xs,rs. + nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → + ∀ls0,y,y0,target,rs0.|xs| = |target| → + nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → + outt = change_vec ?? + (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src) + (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧ + (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → + ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → + outt = int). -lemma sem_rewind : ∀src,dst,sig,n. +lemma sem_rewind_strong : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n → - rewind src dst sig n ⊨ R_rewind src dst sig n. + rewind src dst sig n ⊨ R_rewind_strong src dst sig n. #src #dst #sig #n #Hneq #Hsrc #Hdst @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst) ?) [| @(sem_seq_app sig n ????? (sem_move_multi … R ?) (sem_move_multi … R ?)) // @le_S_S_to_le // ] -#ta #tb * #tc * * #Htc #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb % +#ta #tb * #tc * * * #Htc1 #Htc2 #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb % [ % [ % [ #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst - >(Htc ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd; + >(Htc1 ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd; [|>Hmidta_dst // |>length_append >length_append >Hlen % ] >change_vec_commute [|@sym_not_eq //] @@ -156,8 +72,20 @@ lemma sem_rewind : ∀src,dst,sig,n. >rev_append_def >append_nil #Htd >Htd in Htb; >change_vec_change_vec >nth_change_vec // cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); // +| #x #x0 #xs #rs #Hmidta_dst #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_src + >(Htc2 ??? Hmidta_dst ls0 y (target@[y0]) rs0 ??) in Htd; + [|>Hmidta_src // + |>length_append >length_append >Hlen % ] + >change_vec_change_vec + >change_vec_commute [|@sym_not_eq //] + >nth_change_vec // + >reverse_append >reverse_single + >reverse_append >reverse_single + cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); + #Htd >Htd in Htb; >change_vec_change_vec >nth_change_vec // + >rev_append_def >change_vec_commute // normalize in match (tape_move ???); // ] | #x #rs #Hmidta_src #ls0 #y #rs0 #Hmidta_dst - lapply (Htc … Hmidta_src … (refl ??) Hmidta_dst) -Htc #Htc >Htc in Htd; + lapply (Htc1 … Hmidta_src … (refl ??) Hmidta_dst) -Htc1 #Htc >Htc in Htd; >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%); @@ -169,14 +97,42 @@ lemma sem_rewind : ∀src,dst,sig,n. >nth_change_vec // >change_vec_change_vec whd in match (tape_move ???);whd in match (tape_move ???); change_vec_same >change_vec_same // -]] + ] ] +| #x #rs #Hmidta_dst #ls0 #y #rs0 #Hmidta_src + lapply (Htc2 … Hmidta_dst … (refl ??) Hmidta_src) -Htc2 #Htc >Htc in Htd; + >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] + >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%); + [ #Hls0 destruct (Hls0) #Htd >Htd in Htb; + >nth_change_vec // >change_vec_change_vec + whd in match (tape_move ???);whd in match (tape_move ???); + change_vec_same >change_vec_same // + | #l1 #ls1 #Hls0 destruct (Hls0) #Htd >Htd in Htb; + >nth_change_vec // >change_vec_change_vec + whd in match (tape_move ???); whd in match (tape_move ???); change_vec_same >change_vec_same // + ] +] +qed. + +lemma sem_rewind : ∀src,dst,sig,n. + src ≠ dst → src < S n → dst < S n → + rewind src dst sig n ⊨ R_rewind src dst sig n. +#src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_rewind_strong …)) // +#ta #tb * * * #H1 #H2 #H3 #H4 % /2 by / qed. +(* match step *) + +definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n. + match (nth src (option sig) v (None ?)) with + [ None ⇒ false + | Some x ⇒ notb (nth dst (DeqOption sig) v (None ?) == None ?) ]. + definition match_step ≝ λsrc,dst,sig,n. compare src dst sig n · (ifTM ?? (partest sig n (match_test src dst sig ?)) (single_finalTM ?? - (rewind src dst sig n · (inject_TM ? (move_r ?) n dst))) + (rewind src dst sig n · mmove dst ?? R)) (nop …) partest1). @@ -237,8 +193,8 @@ lemma sem_match_step : (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?)) (sem_seq … (sem_rewind ???? Hneq Hsrc Hdst) - (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? ))) - (sem_nop …))) + (sem_move_multi … R ?)) + (sem_nop …))) /2/ [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?)))) cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%); [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %] @@ -278,15 +234,10 @@ lemma sem_match_step : lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte >change_vec_commute // >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte - >Hte in Htb; * * #_ >nth_change_vec // #Htb1 - lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 % - [ @(eq_vec … (niltape ?)) #i #Hi - cases (true_or_false ((dst : DeqNat) == i)) #Hdsti - [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst - >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] % - | nth_change_vec_neq [| @(\Pf Hdsti)] - >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%); - change_vec_same % ] + >Hte in Htb; whd in ⊢ (%→?); #Htb >Htb % + [ >change_vec_change_vec >nth_change_vec // + >reverse_reverse change_vec_same Hmidta_dst %{s'} % [%] #_ >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // ] @@ -294,14 +245,10 @@ lemma sem_match_step : | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %] -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ - #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2 + #Htd destruct (Htd) * #te * * #_ #Hte whd in ⊢ (%→?); #Htb #s1 #rs1 >Hmidta_src #H destruct (H) lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) % - [ @(eq_vec … (niltape ?)) #i #Hi - cases (true_or_false ((dst : DeqNat) == i)) #Hdsti - [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst - cases rs0 [|#r2 #rs2] % - | nth_change_vec_neq [| @(\Pf Hdsti)] % ] + [ >Htb % | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ] ] ] @@ -482,19 +429,13 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) / ] qed. -axiom daemon : ∀P:Prop.P. - definition R_match_step_true_naive ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n). |left ? (nth src ? outt (niltape ?))| + - |right ? (nth dst ? outt (niltape ?))| < + |option_cons ? (current ? (nth dst ? outt (niltape ?))) (right ? (nth dst ? outt (niltape ?)))| < |left ? (nth src ? int (niltape ?))| + - |right ? (nth dst ? int (niltape ?))|. + |option_cons ? (current ? (nth dst ? int (niltape ?))) (right ? (nth dst ? int (niltape ?)))|. -axiom right_mk_tape : ∀sig,ls,c,rs.right ? (mk_tape sig ls c rs) = rs. -axiom left_mk_tape : ∀sig,ls,c,rs.right ? (mk_tape sig ls c rs) = rs. -axiom length_tail : ∀A,l.0 < |l| → |tail A l| < |l|. - lemma sem_match_step_termination : ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → match_step src dst sig n ⊨ @@ -505,9 +446,9 @@ lemma sem_match_step_termination : @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst) (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?)) (sem_seq … - (sem_rewind ???? Hneq Hsrc Hdst) - (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? ))) - (sem_nop …))) + (sem_rewind_strong ???? Hneq Hsrc Hdst) + (sem_move_multi … R ?)) + (sem_nop …))) [/2/] [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?)))) cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%); [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %] @@ -529,7 +470,7 @@ lemma sem_match_step_termination : >nth_current_chars >nth_current_chars >Hcurtc_dst cases (current ? (nth src …)) [normalize in ⊢ (%→?); #H destruct (H) - | #x >nth_change_vec // cases (reverse ? rs0) + | #x >nth_change_vec [|@Hdst] cases (reverse ? rs0) [ normalize in ⊢ (%→?); #H destruct (H) | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ] | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?); @@ -540,68 +481,195 @@ lemma sem_match_step_termination : ]] normalize in ⊢ (%→?); #H destruct (H) ] | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0 - #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * * - >Htc >change_vec_commute // >nth_change_vec // - >change_vec_commute [|@sym_not_eq //] >nth_change_vec // #Hte #_ - whd in ⊢ (?→%); >Hmidta_src >Hmidta_dst normalize in ⊢ (?→??%); - lapply Hte -Hte @(list_elim_left … ls) - [ #Hte lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte - >change_vec_commute // >change_vec_change_vec - >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte - >Hte * * #_ >nth_change_vec // >reverse_reverse - #H lapply (H … (refl ??)) -H #Htb1 #Htb2 - cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (midtape sig [] s0 (xs@ci::rs'')) src) (mk_tape sig (s0::ls0) (option_hd sig (xs@cj::rs0')) (tail sig (xs@cj::rs0'))) dst) - [@daemon] -Htb1 -Htb2 #Htb >Htb >nth_change_vec // - >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec // - >right_mk_tape normalize in match (left ??); - >length_tail >Hrs0 >length_append whd in ⊢ (?(?%)?); - >commutative_plus % - | #l1 #ls1 #_ >(?:reverse ? xs@s0::ls1@[l1] = (reverse ? xs@s0::ls1)@[l1]) - [|@daemon] - #Hte lapply (Hte ???? (refl ??) ??? (reverse ? xs@s0::ls1) ???) - STOP - (refl ??)) -Hte - >change_vec_commute // >change_vec_change_vec - >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte - >Hte * * #_ >nth_change_vec // >reverse_reverse - #H lapply (H … (refl ??)) -H #Htb1 #Htb2 - cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (midtape sig [] s0 (xs@ci::rs'')) src) (mk_tape sig (s0::ls0) (option_hd sig (xs@cj::rs0')) (tail sig (xs@cj::rs0'))) dst) - [@daemon] -Htb1 -Htb2 #Htb >Htb >nth_change_vec // - >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec // - >right_mk_tape normalize in match (left ??); - >length_tail >Hrs0 >length_append whd in ⊢ (?(?%)?); - >commutative_plus % - - - la - >Hte in Htb; * * #_ >nth_change_vec // #Htb1 - lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 % - [ @(eq_vec … (niltape ?)) #i #Hi - cases (true_or_false ((dst : DeqNat) == i)) #Hdsti - [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst - >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] % - | nth_change_vec_neq [| @(\Pf Hdsti)] - >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%); - change_vec_same % ] - | >Hmidta_dst %{s'} % [%] #_ - >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // + #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * * * + >Htc >change_vec_commute [|//] >nth_change_vec [|//] + >change_vec_commute [|@sym_not_eq //] >nth_change_vec [|//] + cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb + destruct (Hlsalsb) * + [ #Hte #_ #_ <(reverse_reverse … ls) in Hte; <(reverse_reverse … lsa) + cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen' + @(list_cases2 … Hlen') + [ #H1 #H2 >H1 >H2 -H1 -H2 normalize in match (reverse ? [ ]); #Hte #_ + lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte + >change_vec_commute [|//] >change_vec_change_vec + >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte + >Hte whd in ⊢ (%→?); >change_vec_change_vec >nth_change_vec [|//] + >reverse_reverse #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (midtape sig [ ] s0 (xs@ci::rs'')) src) (mk_tape sig (s0::lsb) (option_hd sig (xs@cj::rs0')) (tail sig (xs@cj::rs0'))) dst) + [ >Htb @eq_f3 // cases (xs@cj::rs0') // ] + -Htb #Htb >Htb whd >nth_change_vec [|//] + >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec [|//] + >right_mk_tape [|cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H)] + normalize in match (left ??); + >Hmidta_src >Hmidta_dst >current_mk_tape Hrs0 + normalize in ⊢ (?(?%)%); // + | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2 + >reverse_cons >reverse_cons #Hte + lapply (Hte ci hdb (reverse ? xs@s0::reverse ? tlb) rs'' ? + lsb cj hda (reverse ? xs@s0::reverse ? tla) rs0' ??) + [ /2 by cons_injective_l, nil/ + | >length_append >length_append @eq_f @(eq_f ?? S) + >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' + >length_reverse >length_reverse destruct (Hlen') // + | /2 by refl, trans_eq/ ] -Hte + #Hte #_ whd in ⊢ (%→?); #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta + (mk_tape sig (hda::lsb) (option_hd ? (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0')) (tail ? (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0'))) dst) + (midtape ? [ ] hdb (reverse sig (reverse sig xs@s0::reverse sig tlb)@ci::rs'')) src) + [ >Htb >Hte >nth_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 [|//] + @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0') // ] + -Htb #Htb >Htb whd + >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//] + >right_mk_tape + [| cases (reverse sig (reverse sig xs@s0::reverse sig tla)) + [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] + >Hmidta_src >Hmidta_dst + whd in match (left ??); whd in match (left ??); whd in match (right ??); + >current_mk_tape Hrs0 >length_append whd in ⊢ (??(??%)); >length_append >length_reverse + >length_append >commutative_plus in match (|reverse ??| + ?); + whd in match (|?::?|); >length_reverse >length_reverse + <(length_reverse ? ls) H1 normalize // ] + | #_ #Hte #_ <(reverse_reverse … ls0) in Hte; <(reverse_reverse … lsa) + cut (|reverse ? lsa| = |reverse ? ls0|) [ // ] #Hlen' + @(list_cases2 … Hlen') + [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #Hte + lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte + >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] + >change_vec_change_vec #Hte #_ + >Hte whd in ⊢ (%→?); >nth_change_vec [|//] >reverse_reverse #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (mk_tape ? [s0] (option_hd ? (xs@cj::rs0')) (tail ? (xs@cj::rs0'))) dst) + (midtape ? lsb s0 (xs@ci::rs'')) src) + [ >Htb >change_vec_change_vec >change_vec_commute [|//] + @eq_f3 // Htb whd >nth_change_vec [|//] + >nth_change_vec_neq [|//] >nth_change_vec [|//] + >right_mk_tape + [| cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] + normalize in match (left ??); + >Hmidta_src >Hmidta_dst >current_mk_tape Hrs0 + >length_append normalize >length_append >length_append + <(reverse_reverse ? lsa) >H1 normalize // + | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2 + >reverse_cons >reverse_cons #Hte + lapply (Hte cj hdb (reverse ? xs@s0::reverse ? tlb) rs0' ? + lsb ci hda (reverse ? xs@s0::reverse ? tla) rs'' ??) + [ /2 by cons_injective_l, nil/ + | >length_append >length_append @eq_f @(eq_f ?? S) + >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' + >length_reverse >length_reverse destruct (Hlen') // + | /2 by refl, trans_eq/ ] -Hte + #Hte #_ whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta + (mk_tape sig [hdb] (option_hd ? (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0')) (tail ? (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0'))) dst) + (midtape ? lsb hda (reverse sig (reverse sig xs@s0::reverse sig tla)@ci::rs'')) src) + [ >Htb >change_vec_change_vec >change_vec_commute [|//] + >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] + >change_vec_change_vec >change_vec_commute [|//] + @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0') // ] + -Htb #Htb >Htb whd + >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//] + >right_mk_tape + [| cases (reverse sig (reverse sig xs@s0::reverse sig tlb)) + [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] + >Hmidta_src >Hmidta_dst + whd in match (left ??); whd in match (left ??); whd in match (right ??); + >current_mk_tape Hrs0 >length_append whd in ⊢ (??(??%)); >length_append >length_reverse + >length_append >commutative_plus in match (|reverse ??| + ?); + whd in match (|?::?|); >length_reverse >length_reverse + <(length_reverse ? lsa) >Hlen' >H2 >length_append + normalize // + ] ] ] | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %] -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ - #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2 - #s1 #rs1 >Hmidta_src #H destruct (H) - lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) % - [ @(eq_vec … (niltape ?)) #i #Hi - cases (true_or_false ((dst : DeqNat) == i)) #Hdsti - [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst - cases rs0 [|#r2 #rs2] % - | nth_change_vec_neq [| @(\Pf Hdsti)] % ] - | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ] - ] - ] - ] + #Htd destruct (Htd) * #te * * * * >Hmidta_src >Hmidta_dst + cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb + destruct (Hlsalsb) + [ <(reverse_reverse … ls) <(reverse_reverse … lsa) + cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen' + @(list_cases2 … Hlen') + [ #H1 #H2 >H1 >H2 -H1 -H2 #_ #_ normalize in match (reverse ? [ ]); #Hte #_ + lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte) + whd in ⊢ (%→?); >Hmidta_dst #Htb + cut (tb = change_vec ?? ta (mk_tape ? (s0::lsa@lsb) (option_hd ? rs0) (tail ? rs0)) dst) + [ >Htb cases rs0 // ] + -Htb #Htb >Htb whd >nth_change_vec [|//] + >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst + >right_mk_tape + [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H)] ] + normalize in match (left ??); normalize in match (right ??); + >Hmidta_src >Hmidta_dst >current_mk_tape H1 >H2 + >reverse_cons >reverse_cons >associative_append #Hte + lapply (Hte ???? (refl ??) ? s0 ? (reverse ? tla) ?? (refl ??)) + [ >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) // + normalize #H destruct (H) // ] #Hte #_ #_ #_ + whd in ⊢ (%→?); >Hte >change_vec_change_vec >nth_change_vec // #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta + (mk_tape sig (hda::lsb) (option_hd ? (reverse sig (reverse sig tla)@s0::rs0)) (tail ? (reverse sig (reverse sig tla)@s0::rs0))) dst) + (midtape ? [ ] hdb (reverse sig (reverse sig tlb)@s::rs)) src) + [ >Htb >change_vec_commute [|//] @eq_f3 // cases (reverse sig (reverse sig tla)@s0::rs0) // ] + -Htb #Htb >Htb whd + >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//] + >right_mk_tape + [| cases (reverse sig (reverse sig tla)) + [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] + >Hmidta_src >Hmidta_dst + whd in match (left ??); whd in match (left ??); whd in match (right ??); + >current_mk_tape length_append + >length_reverse >length_reverse <(length_reverse ? ls) H1 normalize // ] + | #_ <(reverse_reverse … ls0) <(reverse_reverse … lsa) + cut (|reverse ? lsa| = |reverse ? ls0|) [ // ] #Hlen' + @(list_cases2 … Hlen') + [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #_ #_ #Hte + lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte) + whd in ⊢ (%→?); #Htb whd >Hmidta_dst + cut (tb = change_vec (tape sig) (S n) ta (mk_tape ? (s0::ls0) (option_hd ? rs0) (tail ? rs0)) dst) + [ >Htb >Hmidta_dst cases rs0 // ] + -Htb #Htb >Htb whd >nth_change_vec [|//] + >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst + >current_mk_tape >right_mk_tape + [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H) ]] + normalize in ⊢ (??%); H1 >H2 + >reverse_cons >reverse_cons #Hte #_ #_ + lapply (Hte s0 hdb (reverse ? tlb) rs0 ? + lsb s hda (reverse ? tla) rs ??) + [ /2 by cons_injective_l, nil/ + | >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) // + normalize #H destruct (H) // + | /2 by refl, trans_eq/ ] -Hte + #Hte whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb + cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta + (mk_tape sig [hdb] (option_hd ? (reverse sig (reverse sig tlb)@s0::rs0)) (tail ? (reverse sig (reverse sig tlb)@s0::rs0))) dst) + (midtape ? lsb hda (reverse sig (reverse sig tla)@s::rs)) src) + [ >Htb >change_vec_commute [|//] >change_vec_change_vec + @eq_f3 // cases (reverse sig (reverse sig tlb)@s0::rs0) // ] + -Htb #Htb >Htb whd + >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//] + >right_mk_tape + [| cases (reverse ? (reverse ? tlb)) [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] + >Hmidta_src >Hmidta_dst + whd in match (left ??); whd in match (left ??); whd in match (right ??); + >current_mk_tape length_append + normalize in ⊢ (??%); >length_append >reverse_reverse + <(length_reverse ? lsa) >Hlen' >H2 normalize // + ] + ] + ] + ] + ] | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd) whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src lapply (refl ? (current ? (nth dst ? ta (niltape ?)))) @@ -631,50 +699,31 @@ lemma sem_match_step_termination : normalize #H destruct (H) ] ] ] qed. +(* lemma WF_to_WF_f : ∀A,B,R,f,b. WF A R (f b) → WF B (λx,y.R (f x) (f y)) b. *) +let rec WF_to_WF_f A B R f b (Hwf: WF A R (f b)) on Hwf: WF B (λx,y.R (f x) (f y)) b ≝ + match Hwf return (λa0,r.f b = a0 → WF B (λx,y:B. R (f x) (f y)) b) with + [ wf a Hwfa ⇒ λHeq.? ] (refl ??). +% #b1 #HRb @WF_to_WF_f @Hwfa nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ * #s0 * normalize in ⊢ (%→?); #H destruct (H) -| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?); - * #s0 * #H destruct (H) -| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?); - * #s0 * #H destruct (H) -| #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs - [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ?? (refl ??)) -HR - >change_vec_change_vec #Ht1 #_ % #t2 whd in ⊢ (%→?); - >Ht1 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src #HR - cases (HR ?? (refl ??)) -HR #_ - >nth_change_vec // * #s1 * normalize in ⊢ (%→?); #H destruct (H) - |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); - >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src - #Htrue cases (Htrue ?? (refl ??)) -Htrue >change_vec_change_vec - >nth_change_vec // >Hmid_dst whd in match (tape_move_mono ???); #Ht1 - * #s0 * whd in ⊢ (??%?→?); #H destruct (H) #_ >Ht1 - lapply (IH t1 ? (s0::ls) r0 ?) - [ >Ht1 >nth_change_vec // - | >Ht1 >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src - | >Ht1 >nth_change_vec // ] - ] -] +#src #dst #sig #n #ta #Hneq #Hsrc #Hdst +@(terminate_while … (sem_match_step_termination src dst sig n Hneq Hsrc Hdst)) // +letin f ≝ (λt0:Vector (tape sig) (S n).|left ? (nth src (tape ?) t0 (niltape ?))| + +|option_cons ? (current ? (nth dst (tape ?) t0 (niltape ?))) + (right ? (nth dst (tape ?) t0 (niltape ?)))|) +change with (λx,y.f x < f y) in ⊢ (??%?); @WF_to_WF_f @lt_WF +qed. + +lemma sem_match_m : ∀src,dst,sig,n. +src ≠ dst → src < S n → dst < S n → + match_m src dst sig n \vDash R_match_m src dst sig n. +#src #dst #sig #n #Hneq #Hsrc #Hdst @WRealize_to_Realize [/2/| @wsem_match_m // ] qed. \ No newline at end of file