X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Fmarks.ma;h=bfa8b8a748b48c7702770253ac1809da712f8971;hb=ce60ad8a7d4c56f218d95c3547abe896057de040;hp=17992f95fca7cdef66e89e4f82c85abc174f338b;hpb=67948c54e4e20ef530d2d2115c9a056275360012;p=helm.git diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index 17992f95f..bfa8b8a74 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -14,10 +14,9 @@ *) -include "turing/while_machine.ma". include "turing/if_machine.ma". +include "turing/basic_machines.ma". include "turing/universal/alphabet.ma". -include "turing/universal/tests.ma". (* ADVANCE TO MARK (right) @@ -30,17 +29,21 @@ include "turing/universal/tests.ma". definition atm_states ≝ initN 3. +definition atm0 : atm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition atm1 : atm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition atm2 : atm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition atmr_step ≝ λalpha:FinSet.λtest:alpha→bool. mk_TM alpha atm_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1, None ?〉 + [ None ⇒ 〈atm1, None ?〉 | Some a' ⇒ match test a' with - [ true ⇒ 〈1,None ?〉 - | false ⇒ 〈2,Some ? 〈a',R〉〉 ]]) - O (λx.notb (x == 0)). + [ true ⇒ 〈atm1,None ?〉 + | false ⇒ 〈atm2,Some ? 〈a',R〉〉 ]]) + atm0 (λx.notb (x == atm0)). definition Ratmr_step_true ≝ λalpha,test,t1,t2. @@ -57,43 +60,45 @@ definition Ratmr_step_false ≝ lemma atmr_q0_q1 : ∀alpha,test,ls,a0,rs. test a0 = true → step alpha (atmr_step alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (atmr_step alpha test)) 1 + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atmr_step alpha test)) atm1 (midtape … ls a0 rs). -#alpha #test #ls #a0 #ts #Htest normalize >Htest % +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest % qed. lemma atmr_q0_q2 : ∀alpha,test,ls,a0,rs. test a0 = false → step alpha (atmr_step alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (atmr_step alpha test)) 2 + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atmr_step alpha test)) atm2 (mk_tape … (a0::ls) (option_hd ? rs) (tail ? rs)). -#alpha #test #ls #a0 #ts #Htest normalize >Htest cases ts // +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest cases ts // qed. lemma sem_atmr_step : ∀alpha,test. accRealize alpha (atmr_step alpha test) - 2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test). + atm2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test). #alpha #test * [ @(ex_intro ?? 2) - @(ex_intro ?? (mk_config ?? 1 (niltape ?))) % - [ % // #Hfalse destruct | #_ % // % % ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (leftof ? a al))) - % [ % // #Hfalse destruct | #_ % // % % ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (rightof ? a al))) - % [ % // #Hfalse destruct | #_ % // % % ] + @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) % + [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] | #ls #c #rs @(ex_intro ?? 2) cases (true_or_false (test c)) #Htest - [ @(ex_intro ?? (mk_config ?? 1 ?)) + [ @(ex_intro ?? (mk_config ?? atm1 ?)) [| % [ % [ whd in ⊢ (??%?); >atmr_q0_q1 // - | #Hfalse destruct ] + | whd in ⊢ ((??%%)→?); #Hfalse destruct ] | #_ % // %2 @(ex_intro ?? c) % // ] ] - | @(ex_intro ?? (mk_config ?? 2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs)))) + | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs)))) % [ % [ whd in ⊢ (??%?); >atmr_q0_q2 // @@ -106,17 +111,34 @@ lemma sem_atmr_step : ] qed. +lemma dec_test: ∀alpha,rs,test. + decidable (∀c.memb alpha c rs = true → test c = false). +#alpha #rs #test elim rs + [%1 #n normalize #H destruct + |#a #tl cases (true_or_false (test a)) #Ha + [#_ %2 % #Hall @(absurd ?? not_eq_true_false) (\P eqca) @Ha |@Hall] + |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons // + ] + qed. + definition R_adv_to_mark_r ≝ λalpha,test,t1,t2. + (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,rs. (t1 = midtape alpha ls c rs → ((test c = true ∧ t2 = t1) ∨ (test c = false ∧ - ∀rs1,b,rs2. rs = rs1@b::rs2 → + (∀rs1,b,rs2. rs = rs1@b::rs2 → test b = true → (∀x.memb ? x rs1 = true → test x = false) → - t2 = midtape ? (reverse ? rs1@c::ls) b rs2))). + t2 = midtape ? (reverse ? rs1@c::ls) b rs2) ∧ + ((∀x.memb ? x rs = true → test x = false) → + ∀a,l.reverse ? (c::rs) = a::l → + t2 = rightof alpha a (l@ls))))). definition adv_to_mark_r ≝ - λalpha,test.whileTM alpha (atmr_step alpha test) 2. + λalpha,test.whileTM alpha (atmr_step alpha test) atm2. lemma wsem_adv_to_mark_r : ∀alpha,test. @@ -125,31 +147,55 @@ lemma wsem_adv_to_mark_r : lapply (sem_while … (sem_atmr_step alpha test) t i outc Hloop) [%] -Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar) [ #tapea * #Htapea * - [ #H1 #ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?); - #Hfalse destruct (Hfalse) - | * #a * #Ha #Htest #ls #c #rs #H2 % - >H2 in Ha; whd in ⊢ (??%? → ?); #Heq destruct (Heq) % // - H2 in H1; whd in ⊢ (??%? → ?); + #Hfalse destruct (Hfalse) + ] + | * #a * #Ha #Htest % + [ >Ha #H destruct (H); + | #ls #c #rs #H2 % + >H2 in Ha; whd in ⊢ (??%? → ?); #Heq destruct (Heq) % // + Htapea' in Htapea; #Htapea destruct (Htapea) % // * - [ #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #_ - cases (IH … Htapeb) - [ * #_ #Houtc >Houtc >Htapeb % - | * #Hfalse >Hfalse in Htestb; #Htestb destruct (Htestb) ] - | #r1 #rs1 #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #Hmemb - cases (IH … Htapeb) - [ * #Hfalse >(Hmemb …) in Hfalse; - [ #Hft destruct (Hft) - | @memb_hd ] - | * #Htestr1 #H1 >reverse_cons >associative_append - @H1 // #x #Hx @Hmemb @memb_cons // + lapply (IH HRfalse) -IH #IH % + [cases Hleft #ls * #a * #rs * * #Htapea #_ #_ >Htapea + whd in ⊢((??%?)→?); #H destruct (H); + |#ls #c #rs #Htapea %2 + cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest #Htapeb + >Htapea' in Htapea; #Htapea destruct (Htapea) % [ % // ] + [* + [ #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #_ + cases (proj2 ?? IH … Htapeb) + [ * #_ #Houtc >Houtc >Htapeb % + | * * >Htestb #Hfalse destruct (Hfalse) ] + | #r1 #rs1 #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #Hmemb + cases (proj2 ?? IH … Htapeb) + [ * #Hfalse >(Hmemb …) in Hfalse; + [ #Hft destruct (Hft) + | @memb_hd ] + | * * #Htestr1 #H1 #_ >reverse_cons >associative_append + @H1 // #x #Hx @Hmemb @memb_cons // + ] ] - ] + |cases rs in Htapeb; normalize in ⊢ (%→?); + [#Htapeb #_ #a0 #l whd in ⊢ ((??%?)→?); #Hrev destruct (Hrev) + >Htapeb in IH; #IH cases (proj1 ?? IH … (refl …)) // + |#r1 #rs1 #Htapeb #Hmemb + cases (proj2 ?? IH … Htapeb) [ * >Hmemb [ #Hfalse destruct(Hfalse) ] @memb_hd ] + * #_ #H1 #a #l <(reverse_reverse … l) cases (reverse … l) + [#H cut (c::r1::rs1 = [a]) + [<(reverse_reverse … (c::r1::rs1)) >H //] + #Hrev destruct (Hrev) + |#a1 #l2 >reverse_cons >reverse_cons >reverse_cons + #Hrev cut ([c] = [a1]) + [@(append_l2_injective_r ?? (a::reverse … l2) … Hrev) //] + #Ha associative_append @H1 + [#x #membx @Hmemb @memb_cons @membx + |<(append_l1_injective_r ?? (a::reverse … l2) … Hrev) // + ] qed. lemma terminate_adv_to_mark_r : @@ -188,105 +234,40 @@ qed. definition mark_states ≝ initN 2. +definition ms0 : mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)). +definition ms1 : mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)). + definition mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mark_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',true〉,N〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈ms1,None ?〉 + | Some a' ⇒ match (pi1 … q) with + [ O ⇒ let 〈a'',b〉 ≝ a' in 〈ms1,Some ? 〈〈a'',true〉,N〉〉 + | S q ⇒ 〈ms1,None ?〉 ] ]) + ms0 (λq.q == ms1). definition R_mark ≝ λalpha,t1,t2. - ∀ls,c,b,rs. - t1 = midtape (FinProd … alpha FinBool) ls 〈c,b〉 rs → - t2 = midtape ? ls 〈c,true〉 rs. + (∀ls,c,b,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈c,b〉 rs → + t2 = midtape ? ls 〈c,true〉 rs) ∧ + (current ? t1 = None ? → t2 = t1). lemma sem_mark : ∀alpha.Realize ? (mark alpha) (R_mark alpha). #alpha #intape @(ex_intro ?? 2) cases intape [ @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#ls #c #b #rs #Hfalse destruct | // ]]] |#a #al @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#ls #c #b #rs #Hfalse destruct | // ]]] |#a #al @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#ls #c #b #rs #Hfalse destruct ] // ]] | #ls * #c #b #rs - @ex_intro [| % [ % | #ls0 #c0 #b0 #rs0 #H1 destruct (H1) % ] ] ] + @ex_intro [| % [ % | % + [#ls0 #c0 #b0 #rs0 #H1 destruct (H1) % + | whd in ⊢ ((??%?)→?); #H1 destruct (H1)]]] qed. -(* MOVE RIGHT - - moves the head one step to the right - -*) - -definition move_states ≝ initN 2. - -definition move_r ≝ - λalpha:FinSet.mk_TM alpha move_states - (λp.let 〈q,a〉 ≝ p in - match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',R〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). - -definition R_move_r ≝ λalpha,t1,t2. - ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs). - -lemma sem_move_r : - ∀alpha.Realize ? (move_r alpha) (R_move_r alpha). -#alpha #intape @(ex_intro ?? 2) cases intape -[ @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -| #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases rs0 // ] ] ] -qed. - -(* MOVE LEFT - - moves the head one step to the right - -*) - -definition move_l ≝ - λalpha:FinSet.mk_TM alpha move_states - (λp.let 〈q,a〉 ≝ p in - match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',L〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). - -definition R_move_l ≝ λalpha,t1,t2. - ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (tail ? ls) (option_hd ? ls) (c::rs). - -lemma sem_move_l : - ∀alpha.Realize ? (move_l alpha) (R_move_l alpha). -#alpha #intape @(ex_intro ?? 2) cases intape -[ @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -| #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases ls0 // ] ] ] -qed. (* MOVE RIGHT AND MARK machine @@ -297,18 +278,22 @@ qed. definition mrm_states ≝ initN 3. +definition mrm0 : mrm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition mrm1 : mrm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition mrm2 : mrm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition move_right_and_mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mrm_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈2,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',R〉〉 + [ None ⇒ 〈mrm2,None ?〉 + | Some a' ⇒ match pi1 … q with + [ O ⇒ 〈mrm1,Some ? 〈a',R〉〉 | S q ⇒ match q with [ O ⇒ let 〈a'',b〉 ≝ a' in - 〈2,Some ? 〈〈a'',true〉,N〉〉 - | S _ ⇒ 〈2,None ?〉 ] ] ]) - O (λq.q == 2). + 〈mrm2,Some ? 〈〈a'',true〉,N〉〉 + | S _ ⇒ 〈mrm2,None ?〉 ] ] ]) + mrm0 (λq.q == mrm2). definition R_move_right_and_mark ≝ λalpha,t1,t2. ∀ls,c,d,b,rs. @@ -337,17 +322,22 @@ qed. definition clear_mark_states ≝ initN 3. +definition clear0 : clear_mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition clear1 : clear_mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition claer2 : clear_mark_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition clear_mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) clear_mark_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',false〉,N〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈clear1,None ?〉 + | Some a' ⇒ match pi1 … q with + [ O ⇒ let 〈a'',b〉 ≝ a' in 〈clear1,Some ? 〈〈a'',false〉,N〉〉 + | S q ⇒ 〈clear1,None ?〉 ] ]) + clear0 (λq.q == clear1). definition R_clear_mark ≝ λalpha,t1,t2. + (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,b,rs. t1 = midtape (FinProd … alpha FinBool) ls 〈c,b〉 rs → t2 = midtape ? ls 〈c,false〉 rs. @@ -356,13 +346,14 @@ lemma sem_clear_mark : ∀alpha.Realize ? (clear_mark alpha) (R_clear_mark alpha). #alpha #intape @(ex_intro ?? 2) cases intape [ @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]] |#a #al @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]] |#a #al @ex_intro - [| % [ % | #ls #c #b #rs #Hfalse destruct ] ] + [| % [ % | % [#_ %|#ls #c #b #rs #Hfalse destruct ]]] | #ls * #c #b #rs - @ex_intro [| % [ % | #ls0 #c0 #b0 #rs0 #H1 destruct (H1) % ] ] ] + @ex_intro [| % [ % | % + [whd in ⊢ (??%?→?); #H destruct| #ls0 #c0 #b0 #rs0 #H1 destruct (H1) % ]]]] qed. (* ADVANCE MARK RIGHT machine @@ -374,27 +365,35 @@ qed. definition adv_mark_r ≝ λalpha:FinSet. - seq ? (clear_mark alpha) - (seq ? (move_r ?) (mark alpha)). + clear_mark alpha · move_r ? · mark alpha. definition R_adv_mark_r ≝ λalpha,t1,t2. - ∀ls,c,d,b,rs. - t1 = midtape (FinProd … alpha FinBool) ls 〈c,true〉 (〈d,b〉::rs) → - t2 = midtape ? (〈c,false〉::ls) 〈d,true〉 rs. + (∀ls,c. + (∀d,b,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈c,true〉 (〈d,b〉::rs) → + t2 = midtape ? (〈c,false〉::ls) 〈d,true〉 rs) ∧ + (t1 = midtape (FinProd … alpha FinBool) ls 〈c,true〉 [ ] → + t2 = rightof ? 〈c,false〉 ls)) ∧ + (current ? t1 = None ? → t1 = t2). lemma sem_adv_mark_r : ∀alpha.Realize ? (adv_mark_r alpha) (R_adv_mark_r alpha). -#alpha #intape -cases (sem_seq ????? (sem_clear_mark …) - (sem_seq ????? (sem_move_r ?) (sem_mark alpha)) intape) -#k * #outc * #Hloop whd in ⊢ (%→?); -* #ta * whd in ⊢ (%→?); #Hs1 * #tb * whd in ⊢ (%→%→?); #Hs2 #Hs3 -@(ex_intro ?? k) @(ex_intro ?? outc) % -[ @Hloop -| -Hloop #ls #c #d #b #rs #Hintape @(Hs3 … b) - @(Hs2 ls 〈c,false〉 (〈d,b〉::rs)) - @(Hs1 … Hintape) -] +#alpha +@(sem_seq_app … (sem_clear_mark …) + (sem_seq ????? (sem_move_r ?) (sem_mark alpha)) …) +#intape #outtape whd in ⊢ (%→?); * #ta * +whd in ⊢ (%→?); #Hs1 whd in ⊢ (%→?); * #tb * #Hs2 whd in ⊢ (%→?); #Hs3 % + [#ls #c % + [#d #b #rs #Hintape @(proj1 … Hs3 ?? b ?) + @(proj2 … Hs2 ls 〈c,false〉 (〈d,b〉::rs)) + @(proj2 ?? Hs1 … Hintape) + |#Hintape lapply (proj2 ?? Hs1 … Hintape) #Hta lapply (proj2 … Hs2 … Hta) + whd in ⊢ ((???%)→?); #Htb Htb // + ] + |#Hcur lapply(proj1 ?? Hs1 … Hcur) #Hta >Hta >Hta in Hcur; #Hcur + lapply (proj1 ?? Hs2 … Hcur) #Htb >Htb >Htb in Hcur; #Hcur + @sym_eq @(proj2 ?? Hs3) @Hcur + ] qed. (* ADVANCE TO MARK (left) @@ -402,80 +401,186 @@ qed. axiomatized *) +definition atml_step ≝ + λalpha:FinSet.λtest:alpha→bool. + mk_TM alpha atm_states + (λp.let 〈q,a〉 ≝ p in + match a with + [ None ⇒ 〈atm1, None ?〉 + | Some a' ⇒ + match test a' with + [ true ⇒ 〈atm1,None ?〉 + | false ⇒ 〈atm2,Some ? 〈a',L〉〉 ]]) + atm0 (λx.notb (x == atm0)). + +definition Ratml_step_true ≝ + λalpha,test,t1,t2. + ∃ls,a,rs. + t1 = midtape alpha ls a rs ∧ test a = false ∧ + t2 = mk_tape alpha (tail ? ls) (option_hd ? ls) (a :: rs). + +definition Ratml_step_false ≝ + λalpha,test,t1,t2. + t1 = t2 ∧ + (current alpha t1 = None ? ∨ + (∃a.current ? t1 = Some ? a ∧ test a = true)). + +lemma atml_q0_q1 : + ∀alpha,test,ls,a0,rs. test a0 = true → + step alpha (atml_step alpha test) + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atml_step alpha test)) atm1 + (midtape … ls a0 rs). +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest % +qed. + +lemma atml_q0_q2 : + ∀alpha,test,ls,a0,rs. test a0 = false → + step alpha (atml_step alpha test) + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atml_step alpha test)) atm2 + (mk_tape … (tail ? ls) (option_hd ? ls) (a0 :: rs)). +#alpha #test #ls #a0 #rs #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest cases ls // +qed. + +lemma sem_atml_step : + ∀alpha,test. + accRealize alpha (atml_step alpha test) + atm2 (Ratml_step_true alpha test) (Ratml_step_false alpha test). +#alpha #test * +[ @(ex_intro ?? 2) + @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) % + [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #ls #c #rs @(ex_intro ?? 2) + cases (true_or_false (test c)) #Htest + [ @(ex_intro ?? (mk_config ?? atm1 ?)) + [| % + [ % + [ whd in ⊢ (??%?); >atml_q0_q1 // + | whd in ⊢ ((??%%)→?); #Hfalse destruct ] + | #_ % // %2 @(ex_intro ?? c) % // ] + ] + | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (tail ? ls) (option_hd ? ls) (c::rs)))) + % + [ % + [ whd in ⊢ (??%?); >atml_q0_q2 // + | #_ @(ex_intro ?? ls) @(ex_intro ?? c) @(ex_intro ?? rs) + % // % // + ] + | #Hfalse @False_ind @(absurd ?? Hfalse) % + ] + ] +] +qed. definition R_adv_to_mark_l ≝ λalpha,test,t1,t2. + (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,rs. (t1 = midtape alpha ls c rs → - ((test c = true ∧ t2 = t1) ∨ - (test c = false ∧ - ∀ls1,b,ls2. ls = ls1@b::ls2 → + ((test c = true → t2 = t1) ∧ + (test c = false → + (∀ls1,b,ls2. ls = ls1@b::ls2 → test b = true → (∀x.memb ? x ls1 = true → test x = false) → - t2 = midtape ? ls2 b (reverse ? ls1@c::rs)))). + t2 = midtape ? ls2 b (reverse ? ls1@c::rs)) ∧ + ((∀x.memb ? x ls = true → test x = false) → + ∀a,l. reverse ? (c::ls) = a::l → t2 = leftof ? a (l@rs)) + ))). -axiom adv_to_mark_l : ∀alpha:FinSet.(alpha → bool) → TM alpha. -(* definition adv_to_mark_l ≝ - λalpha,test.whileTM alpha (atml_step alpha test) 2. *) +definition adv_to_mark_l ≝ + λalpha,test.whileTM alpha (atml_step alpha test) atm2. -axiom wsem_adv_to_mark_l : +lemma wsem_adv_to_mark_l : ∀alpha,test. WRealize alpha (adv_to_mark_l alpha test) (R_adv_to_mark_l alpha test). -(* #alpha #test #t #i #outc #Hloop -lapply (sem_while … (sem_atmr_step alpha test) t i outc Hloop) [%] +lapply (sem_while … (sem_atml_step alpha test) t i outc Hloop) [%] -Hloop * #t1 * #Hstar @(star_ind_l ??????? Hstar) [ #tapea * #Htapea * - [ #H1 #ls #c #rs #H2 >H2 in H1; whd in ⊢ (??%? → ?); - #Hfalse destruct (Hfalse) - | * #a * #Ha #Htest #ls #c #rs #H2 % - >H2 in Ha; whd in ⊢ (??%? → ?); #Heq destruct (Heq) % // - H2 in H1; whd in ⊢ (??%? → ?); + #Hfalse destruct (Hfalse) + ] + | * #a * #Ha #Htest % + [>Ha #H destruct (H); + |#ls #c #rs #H2 % + [#Hc H2 in Ha; whd in ⊢ ((??%?)→?); + #H destruct (H) /2/ + ] + ] ] | #tapea #tapeb #tapec #Hleft #Hright #IH #HRfalse - lapply (IH HRfalse) -IH #IH - #ls #c #rs #Htapea %2 - cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest #Htapeb - - >Htapea' in Htapea; #Htapea destruct (Htapea) % // * - [ #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #_ - cases (IH … Htapeb) - [ * #_ #Houtc >Houtc >Htapeb % - | * #Hfalse >Hfalse in Htestb; #Htestb destruct (Htestb) ] - | #r1 #rs1 #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #Hmemb - cases (IH … Htapeb) - [ * #Hfalse >(Hmemb …) in Hfalse; - [ #Hft destruct (Hft) - | @memb_hd ] - | * #Htestr1 #H1 >reverse_cons >associative_append - @H1 // #x #Hx @Hmemb @memb_cons // + lapply (IH HRfalse) -IH #IH % + [cases Hleft #ls0 * #a0 * #rs0 * * #Htapea #_ #_ >Htapea + whd in ⊢ ((??%?)→?); #H destruct (H) + |#ls #c #rs #Htapea % + [#Hc cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest @False_ind + >Htapea' in Htapea; #H destruct /2/ + |cases Hleft #ls0 * #a * #rs0 * + * #Htapea1 >Htapea in Htapea1; #H destruct (H) #_ #Htapeb + #Hc % + [* + [#b #ls2 #Hls >Hls in Htapeb; #Htapeb #Htestb #_ + cases (proj2 ?? IH … Htapeb) #H1 #_ >H1 // >Htapeb % + |#l1 #ls1 #b #ls2 #Hls >Hls in Htapeb; #Htapeb #Htestb #Hmemb + cases (proj2 ?? IH … Htapeb) #_ #H1 >reverse_cons >associative_append + @(proj1 ?? (H1 ?) … (refl …) Htestb …) + [@Hmemb @memb_hd + |#x #memx @Hmemb @memb_cons @memx + ] + ] + |cases ls0 in Htapeb; normalize in ⊢ (%→?); + [#Htapeb #Htest #a0 #l whd in ⊢ ((??%?)→?); #Hrev destruct (Hrev) + >Htapeb in IH; #IH cases (proj1 ?? IH … (refl …)) // + |#l1 #ls1 #Htapeb + cases (proj2 ?? IH … Htapeb) #_ #H1 #Htest #a0 #l + <(reverse_reverse … l) cases (reverse … l) + [#H cut (a::l1::ls1 = [a0]) + [<(reverse_reverse … (a::l1::ls1)) >H //] + #Hrev destruct (Hrev) + |#a1 #l2 >reverse_cons >reverse_cons >reverse_cons + #Hrev cut ([a] = [a1]) + [@(append_l2_injective_r ?? (a0::reverse … l2) … Hrev) //] + #Ha associative_append @(proj2 ?? (H1 ?)) + [@Htest @memb_hd + |#x #membx @Htest @memb_cons @membx + |<(append_l1_injective_r ?? (a0::reverse … l2) … Hrev) // + ] + ] + ] + ] ] ] qed. -*) -axiom terminate_adv_to_mark_l : +lemma terminate_adv_to_mark_l : ∀alpha,test. ∀t.Terminate alpha (adv_to_mark_l alpha test) t. -(* #alpha #test #t -@(terminate_while … (sem_atmr_step alpha test)) +@(terminate_while … (sem_atml_step alpha test)) [ % | cases t [ % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse) |2,3: #a0 #al0 % #t1 * #ls0 * #c0 * #rs0 * * #Hfalse destruct (Hfalse) - | #ls #c #rs generalize in match c; -c generalize in match ls; -ls - elim rs - [#ls #c % #t1 * #ls0 * #c0 * #rs0 * * + | #ls elim ls + [#c #rs % #t1 * #ls0 * #c0 * #rs0 * * #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1; % #t2 * #ls1 * #c1 * #rs1 * * >Ht1 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) - | #r0 #rs0 #IH #ls #c % #t1 * #ls0 * #c0 * #rs0 * * + | #rs0 #r0 #IH #ls #c % #t1 * #ls0 * #c0 * #rs0 * * #H1 destruct (H1) #Hc0 #Ht1 normalize in Ht1; >Ht1 @IH ] ] ] qed. -*) lemma sem_adv_to_mark_l : ∀alpha,test. @@ -498,22 +603,58 @@ qed. ^ *) -definition is_marked ≝ - λalpha.λp:FinProd … alpha FinBool. - let 〈x,b〉 ≝ p in b. +definition adv_both_marks ≝ λalpha. + adv_mark_r alpha · move_l ? · + adv_to_mark_l (FinProd alpha FinBool) (is_marked alpha) · + adv_mark_r alpha. + +definition R_adv_both_marks ≝ λalpha,t1,t2. + ∀ls,x0,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs → + (rs = [ ] → (* first case: rs empty *) + t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧ + (∀l0,x,a,a0,b,l1,l1',l2. + ls = (l1@〈x,true〉::l0) → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = (〈a0,b〉::l2) → + reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' → + t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)). -definition adv_both_marks ≝ - λalpha.seq ? (adv_mark_r alpha) - (seq ? (move_l ?) - (seq ? (adv_to_mark_l (FinProd alpha FinBool) (is_marked alpha)) - (adv_mark_r alpha))). +lemma sem_adv_both_marks : + ∀alpha.Realize ? (adv_both_marks alpha) (R_adv_both_marks alpha). +#alpha +@(sem_seq_app … (sem_adv_mark_r …) + (sem_seq ????? (sem_move_l …) + (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?)) + (sem_adv_mark_r alpha))) …) +#intape #outtape * #tapea * #Hta * #tb * #Htb * #tc * #Htc #Hout +#ls #c #rs #Hintape % + [#Hrs >Hrs in Hintape; #Hintape lapply (proj2 ?? (proj1 ?? Hta … ) … Hintape) -Hta #Hta + lapply (proj1 … Htb) >Hta -Htb #Htb lapply (Htb (refl …)) -Htb #Htb + lapply (proj1 ?? Htc) Htc @(proj2 ?? Hout …) Hrs in Hintape; >Hls #Hintape + @(proj1 ?? (proj1 ?? Hout … ) ? false) -Hout + lapply (proj1 … (proj1 … Hta …) … Hintape) #Htapea + lapply (proj2 … Htb … Htapea) -Htb + whd in match (mk_tape ????) ; #Htapeb + lapply (proj1 ?? (proj2 ?? (proj2 ?? Htc … Htapeb) (refl …))) -Htc #Htc + change with ((?::?)@?) in match (cons ???); reverse_cons + >associative_append @Htc [%|%|@Hmarks] + ] +qed. +(* definition R_adv_both_marks ≝ λalpha,t1,t2. ∀l0,x,a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) → - t1 = midtape (FinProd … alpha FinBool) + (t1 = midtape (FinProd … alpha FinBool) (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 (〈a0,false〉::l2) → - t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (reverse ? l1@〈x0,false〉::〈a0,true〉::l2). + t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (reverse ? l1@〈x0,false〉::〈a0,true〉::l2)) ∧ + (t1 = midtape (FinProd … alpha FinBool) + (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 [] → + t2 = rightof ? 〈x0,false〉 (l1@〈a,false〉::〈x,true〉::l0)). lemma sem_adv_both_marks : ∀alpha.Realize ? (adv_both_marks alpha) (R_adv_both_marks alpha). @@ -530,7 +671,7 @@ cases (sem_seq ????? (sem_adv_mark_r …) | -Hloop #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape @(Hs4 … false) -Hs4 lapply (Hs1 … Hintape) #Hta - lapply (Hs2 … Hta) #Htb + lapply (proj2 … Hs2 … Hta) #Htb cases (Hs3 … Htb) [ * #Hfalse normalize in Hfalse; destruct (Hfalse) | * #_ -Hs3 #Hs3 @@ -542,7 +683,7 @@ cases (sem_seq ????? (sem_adv_mark_r …) | >associative_append % | >reverse_append #Htc @Htc ] ] -qed. +qed. *) (* MATCH AND ADVANCE(f) @@ -564,13 +705,52 @@ definition match_and_adv ≝ definition R_match_and_adv ≝ λalpha,f,t1,t2. - ∀l0,x,a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) → - t1 = midtape (FinProd … alpha FinBool) - (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 (〈a0,false〉::l2) → - (f 〈x0,true〉 = true ∧ t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (reverse ? l1@〈x0,false〉::〈a0,true〉::l2)) - ∨ (f 〈x0,true〉 = false ∧ - t2 = midtape ? (l1@〈a,false〉::〈x,true〉::l0) 〈x0,false〉 (〈a0,false〉::l2)). + ∀ls,x0,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs → + ((* first case: (f 〈x0,true〉 = false) *) + (f 〈x0,true〉 = false) → + t2 = midtape (FinProd … alpha FinBool) ls 〈x0,false〉 rs) ∧ + ((f 〈x0,true〉 = true) → rs = [ ] → (* second case: rs empty *) + t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧ + ((f 〈x0,true〉 = true) → + ∀l0,x,a,a0,b,l1,l1',l2. + (* third case: we expect to have a mark on the left! *) + ls = (l1@〈x,true〉::l0) → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = 〈a0,b〉::l2 → + reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' → + t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)). + +lemma sem_match_and_adv : + ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv alpha f). +#alpha #f #intape +cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape) +#k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc) +% [ @Hloop ] -Hloop +(* +@(sem_if_app … (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?)) +#intape #outape #Htb * #H *) +cases Hif -Hif +[ * #ta * whd in ⊢ (%→%→?); * * #c * #Hcurrent #fc #Hta #Houtc + #ls #x #rs #Hintape >Hintape in Hcurrent; whd in ⊢ ((??%?)→?); #H destruct (H) % + [%[>fc #H destruct (H) + |#_ #Hrs >Hrs in Hintape; #Hintape >Hintape in Hta; #Hta + cases (Houtc … Hta) #Houtc #_ @Houtc // + ] + |#l0 #x0 #a #a0 #b #l1 #l1' #l2 #Hls #Hmarks #Hrs #Hrev >Hintape in Hta; #Hta + @(proj2 ?? (Houtc … Hta) … Hls Hmarks Hrs Hrev) + ] +| * #ta * * #H #Hta * #_ #Houtc #ls #c #rs #Hintape + >Hintape in H; #H lapply(H … (refl …)) #fc % + [%[#_ >Hintape in Hta; #Hta @(Houtc … Hta) + |>fc #H destruct (H) + ] + |>fc #H destruct (H) + ] +] +qed. +(* lemma sem_match_and_adv : ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv alpha f). #alpha #f #intape @@ -579,15 +759,42 @@ cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_ma % [ @Hloop ] -Hloop cases Hif [ * #ta * whd in ⊢ (%→%→?); #Hta #Houtc - #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape - >Hintape in Hta; #Hta cases (Hta … (refl ??)) -Hta #Hf #Hta % % - [ @Hf | @Houtc [ @Hl1 | @Hta ] ] + #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape >Hintape in Hta; + * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hf #Hta % % + [ @Hf | >append_cons >append_cons in Hta; #Hta @(proj1 ?? (Houtc …) …Hta) + [ #x #memx cases (memb_append …memx) + [@Hl1 | -memx #memx >(memb_single … memx) %] + |>reverse_cons >reverse_append % ] ] | * #ta * whd in ⊢ (%→%→?); #Hta #Houtc - #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape - >Hintape in Hta; #Hta cases (Hta … (refl ??)) -Hta #Hf #Hta %2 % - [ @Hf | >(Houtc … Hta) % ] + #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape >Hintape in Hta; + * #Hf #Hta %2 % [ @Hf % | >(proj2 ?? Houtc … Hta) % ] ] qed. +*) + +definition R_match_and_adv_of ≝ + λalpha,t1,t2.current (FinProd … alpha FinBool) t1 = None ? → t2 = t1. + +lemma sem_match_and_adv_of : + ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv_of alpha). +#alpha #f #intape +cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape) +#k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc) +% [ @Hloop ] -Hloop +cases Hif +[ * #ta * whd in ⊢ (%→%→?); #Hta #Houtc #Hcur + cases Hta * #x >Hcur * #Hfalse destruct (Hfalse) +| * #ta * whd in ⊢ (%→%→?); * #_ #Hta * #Houtc #_ (Houtc Hcur) % ] +qed. + +lemma sem_match_and_adv_full : + ∀alpha,f.Realize ? (match_and_adv alpha f) + (R_match_and_adv alpha f ∩ R_match_and_adv_of alpha). +#alpha #f #intape cases (sem_match_and_adv ? f intape) +#i * #outc * #Hloop #HR1 %{i} %{outc} % // % // +cases (sem_match_and_adv_of ? f intape) #i0 * #outc0 * #Hloop0 #HR2 +>(loop_eq … Hloop Hloop0) // +qed. (* if x = c @@ -601,25 +808,48 @@ qed. else M *) -definition comp_step_subcase ≝ - λalpha,c,elseM.ifTM ? (test_char ? (λx.x == c)) - (seq ? (move_r …) - (seq ? (adv_to_mark_r ? (is_marked alpha)) - (match_and_adv ? (λx.x == c)))) +definition comp_step_subcase ≝ λalpha,c,elseM. + ifTM ? (test_char ? (λx.x == c)) + (move_r … · adv_to_mark_r ? (is_marked alpha) · match_and_adv ? (λx.x == c)) elseM tc_true. definition R_comp_step_subcase ≝ λalpha,c,RelseM,t1,t2. - ∀l0,x,rs.t1 = midtape (FinProd … alpha FinBool) l0 〈x,true〉 rs → - (〈x,true〉 = c ∧ - ∀a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) → - rs = 〈a,false〉::l1@〈x0,true〉::〈a0,false〉::l2 → - ((x = x0 ∧ - t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1@〈x0,false〉::〈a0,true〉::l2)) ∨ - (x ≠ x0 ∧ - t2 = midtape (FinProd … alpha FinBool) - (reverse ? l1@〈a,false〉::〈x,true〉::l0) 〈x0,false〉 (〈a0,false〉::l2)))) ∨ - (〈x,true〉 ≠ c ∧ RelseM t1 t2). + ∀ls,x,rs.t1 = midtape (FinProd … alpha FinBool) ls 〈x,true〉 rs → + (〈x,true〉 = c → + ((* test true but no marks in rs *) + (∀c.memb ? c rs = true → is_marked ? c = false) → + ∀a,l. (a::l) = reverse ? (〈x,true〉::rs) → + t2 = rightof (FinProd alpha FinBool) a (l@ls)) ∧ + ∀l1,x0,l2. + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = l1@〈x0,true〉::l2 → + (x = x0 → + l2 = [ ] → (* test true but l2 is empty *) + t2 = rightof ? 〈x0,false〉 ((reverse ? l1)@〈x,true〉::ls)) ∧ + (x = x0 → + ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *) + 〈a,false〉::l1' = l1@[〈x0,false〉] → + l2 = 〈a0,b〉::l2' → + t2 = midtape ? (〈x,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧ + (x ≠ x0 →(* test false *) + t2 = midtape (FinProd … alpha FinBool) ((reverse ? l1)@〈x,true〉::ls) 〈x0,false〉 l2)) ∧ + (〈x,true〉 ≠ c → RelseM t1 t2). + +lemma dec_marked: ∀alpha,rs. + decidable (∀c.memb ? c rs = true → is_marked alpha c = false). +#alpha #rs elim rs + [%1 #n normalize #H destruct + |#a #tl cases (true_or_false (is_marked ? a)) #Ha + [#_ %2 % #Hall @(absurd ?? not_eq_true_false) (\P eqca) @Ha |@Hall] + |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons // + ] + qed. + +(* axiom daemon:∀P:Prop.P. *) lemma sem_comp_step_subcase : ∀alpha,c,elseM,RelseM. @@ -631,33 +861,101 @@ cases (sem_if ? (test_char ? (λx.x == c)) … tc_true (sem_test_char ? (λx.x == c)) (sem_seq ????? (sem_move_r …) (sem_seq ????? (sem_adv_to_mark_r ? (is_marked alpha)) - (sem_match_and_adv ? (λx.x == c)))) Helse intape) + (sem_match_and_adv_full ? (λx.x == c)))) Helse intape) #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop cases HR -HR -[ * #ta * whd in ⊢ (%→?); #Hta * #tb * whd in ⊢ (%→?); #Htb - * #tc * whd in ⊢ (%→?); #Htc whd in ⊢ (%→?); #Houtc - #l0 #x #rs #Hintape cases (true_or_false (〈x,true〉==c)) #Hc - [ % % [ @(\P Hc) ] - #a #l1 #x0 #a0 #l2 #Hl1 #Hrs >Hrs in Hintape; #Hintape - >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta - #Hx #Hta lapply (Htb … Hta) -Htb #Htb - cases (Htc … Htb) [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] - -Htc * #_ #Htc lapply (Htc l1 〈x0,true〉 (〈a0,false〉::l2) (refl ??) (refl ??) Hl1) - -Htc #Htc cases (Houtc ???????? Htc) -Houtc - [ * #Hx0 #Houtc % - % [ <(\P Hx0) in Hx; #Hx lapply (\P Hx) #Hx' destruct (Hx') % - | >Houtc >reverse_reverse % ] - | * #Hx0 #Houtc %2 - % [ <(\P Hx) in Hx0; #Hx0 lapply (\Pf Hx0) @not_to_not #Hx' >Hx' % - | >Houtc % ] - | (* members of lists are invariant under reverse *) @daemon ] - | %2 % [ @(\Pf Hc) ] - >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta #Hx #Hta - >Hx in Hc;#Hc destruct (Hc) ] -| * #ta * whd in ⊢ (%→?); #Hta #Helse #ls #c0 #rs #Hintape %2 - >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta #Hc #Hta % - [ @(\Pf Hc) | Hintape in Hcin; whd in ⊢ ((??%?)→?); #H destruct (H) % + [#_ cases (dec_marked ? rs) #Hdec + [% + [#_ #a #l1 + >Hintape in Hta; #Hta + lapply (proj2 ?? Htb … Hta) -Htb -Hta cases rs in Hdec; + (* by cases on rs *) + [#_ whd in ⊢ ((???%)→?); #Htb >Htb in Htc; #Htc + lapply (proj1 ?? Htc (refl …)) -Htc #Htc Htb in Htc; #Htc + >reverse_cons >reverse_cons #Hl1 + cases (proj2 ?? Htc … (refl …)) + [* >(Hdec …) [ #Hfalse destruct(Hfalse) ] @memb_hd + |* #_ -Htc #Htc cut (∃l2.l1 = l2@[〈x,true〉]) + [generalize in match Hl1; -Hl1 <(reverse_reverse … l1) + cases (reverse ? l1) + [#Hl1 cut ([a]=〈x,true〉::r0::rs0) + [ <(reverse_reverse … (〈x,true〉::r0::rs0)) + >reverse_cons >reverse_cons reverse_cons #Heq + lapply (append_l2_injective_r ? (a::reverse ? l10) ???? Heq) // + #Ha0 destruct(Ha0) /2/ ] + |* #l2 #Hl2 >Hl2 in Hl1; #Hl1 + lapply (append_l1_injective_r ? (a::l2) … Hl1) // -Hl1 #Hl1 + >reverse_cons in Htc; #Htc lapply (Htc … (sym_eq … Hl1)) + [ #x0 #Hmemb @Hdec @memb_cons @Hmemb ] + -Htc #Htc >Htc in Houtc1; #Houtc1 >associative_append @Houtc1 % + ] + ] + ] + |#l1 #x0 #l2 #_ #Hrs @False_ind + @(absurd ?? not_eq_true_false) + change with (is_marked ? 〈x0,true〉) in match true; + @Hdec >Hrs @memb_append_l2 @memb_hd + ] + |% [#H @False_ind @(absurd …H Hdec)] + (* by cases on l1 *) * + [#x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape + >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) + #Hta lapply (proj2 … Htb … Hta) -Htb -Hta + whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc + #_ #Htc cases (Htc … Htb) -Htc + [2: * * #Hfalse normalize in Hfalse; destruct (Hfalse) ] + * * #Htc >Htb in Htc; -Htb #Htc cases (Houtc … Htc) -Houtc * + #H1 #H2 #H3 cases (true_or_false (x==x0)) #eqxx0 + [>(\P eqxx0) % [2: #H @False_ind /2/] % + [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0) [% | @Hcintrue] + |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes destruct (Hdes) + #Hl2 @(H3 … Hdec … Hl2) <(\P eqxx0) [@Hcintrue | % | @reverse_single] + ] + |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)] + #_ @H1 @(\bf ?) @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue) + #Hdes destruct (Hdes) % + ] + |#l1hd #l1tl #x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape + >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) + #Hta lapply (proj2 … Htb … Hta) -Htb -Hta + whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc + #_ #Htc cases (Htc … Htb) -Htc + [* #Hfalse @False_ind >(Hdec … (memb_hd …)) in Hfalse; #H destruct] + * * #_ #Htc lapply (Htc … (refl …) (refl …) ?) -Htc + [#x1 #membx1 @Hdec @memb_cons @membx1] #Htc + cases (Houtc … Htc) -Houtc * + #H1 #H2 #H3 #_ cases (true_or_false (x==x0)) #eqxx0 + [>(\P eqxx0) % [2: #H @False_ind /2/] % + [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0) + [>reverse_cons >associative_append % | @Hcintrue] + |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes (* destruct (Hdes) *) + #Hl2 @(H3 ?????? (reverse … (l1hd::l1tl)) … Hl2) <(\P eqxx0) + [@Hcintrue + |>reverse_cons >associative_append % + |#c0 #memc @Hdec <(reverse_reverse ? (l1hd::l1tl)) @memb_reverse @memc + |>Hdes >reverse_cons >reverse_reverse >(\P eqxx0) % + ] + ] + |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)] + #_ >reverse_cons >associative_append @H1 @(\bf ?) + @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue) #Hdes + destruct (Hdes) % + ] + ] + ] + |>(\P Hcintrue) * #Hfalse @False_ind @Hfalse % + ] + | * #ta * * #Hcur #Hta #Houtc + #l0 #x #rs #Hintape >Hintape in Hcur; #Hcur lapply (Hcur ? (refl …)) -Hcur #Hc % + [ #Hfalse >Hfalse in Hc; #Hc cases (\Pf Hc) #Hc @False_ind @Hc % + | -Hc #Hc Hfa normalize in ⊢ (%→?); + #H destruct + [% [@(ex_intro … []) % normalize [% % | #x @False_ind] + |#a1 #tl1 #H destruct (H) //] + |cases (Hind (a::acc) res1 res2 H) * #l1 * * + #Hres1 #Htl #Hfalse #Htrue % [2:@Htrue] @(ex_intro … (l1@[a])) % + [% [>associative_append @Hres1 | >reverse_append H //| @False_ind] + ] + ] + ] +qed. + +axiom mem_reverse: ∀A,l,x. mem A x (reverse ? l) → mem A x l. + +lemma split_on_spec_ex: ∀A,l,f.∃l1,l2. + l1@l2 = l ∧ (∀x:A. mem ? x l1 → f x = false) ∧ + ∀a,tl. l2 = a::tl → f a = true. +#A #l #f @(ex_intro … (reverse … (\fst (split_on A l f [])))) +@(ex_intro … (\snd (split_on A l f []))) +cases (split_on_spec A l f [ ] ?? (eq_pair_fst_snd …)) * #l1 * * +>append_nil #Hl1 >Hl1 #Hl #Hfalse #Htrue % + [% [@Hl|#x #memx @Hfalse @mem_reverse //] | @Htrue] +qed. + +(* versione esistenziale *) + +definition R_comp_step_true ≝ λt1,t2. + ∃ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls 〈c,true〉 rs ∧ + ((* bit_or_null c = false *) + (bit_or_null c = false → t2 = midtape ? ls 〈c,false〉 rs) ∧ + (* no marks in rs *) + (bit_or_null c = true → + (∀c.memb ? c rs = true → is_marked ? c = false) → + ∀a,l. (a::l) = reverse ? (〈c,true〉::rs) → + t2 = rightof (FinProd FSUnialpha FinBool) a (l@ls)) ∧ + (∀l1,c0,l2. + bit_or_null c = true → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = l1@〈c0,true〉::l2 → + (c = c0 → + l2 = [ ] → (* test true but l2 is empty *) + t2 = rightof ? 〈c0,false〉 ((reverse ? l1)@〈c,true〉::ls)) ∧ + (c = c0 → + ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *) + 〈a,false〉::l1' = l1@[〈c0,false〉] → + l2 = 〈a0,b〉::l2' → + t2 = midtape ? (〈c,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧ + (c ≠ c0 →(* test false *) + t2 = midtape (FinProd … FSUnialpha FinBool) + ((reverse ? l1)@〈c,true〉::ls) 〈c0,false〉 l2))). + +definition R_comp_step_false ≝ + λt1,t2. + ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → + is_marked ? c = false ∧ t2 = t1. + +lemma is_marked_to_exists: ∀alpha,c. is_marked alpha c = true → + ∃c'. c = 〈c',true〉. +#alpha * #c * [#_ @(ex_intro … c) //| normalize #H destruct] +qed. + +lemma exists_current: ∀alpha,c,t. + current alpha t = Some alpha c → ∃ls,rs. t= midtape ? ls c rs. +#alpha #c * + [whd in ⊢ (??%?→?); #H destruct + |#a #l whd in ⊢ (??%?→?); #H destruct + |#a #l whd in ⊢ (??%?→?); #H destruct + |#ls #c1 #rs whd in ⊢ (??%?→?); #H destruct + @(ex_intro … ls) @(ex_intro … rs) // + ] +qed. + +lemma sem_comp_step : + accRealize ? comp_step (inr … (inl … (inr … start_nop))) + R_comp_step_true R_comp_step_false. +@(acc_sem_if_app … (sem_test_char ? (is_marked ?)) + (sem_comp_step_subcase FSUnialpha 〈bit false,true〉 ?? + (sem_comp_step_subcase FSUnialpha 〈bit true,true〉 ?? + (sem_comp_step_subcase FSUnialpha 〈null,true〉 ?? + (sem_clear_mark …)))) + (sem_nop …) …) +[#intape #outape #ta #Hta #Htb cases Hta * #cm * #Hcur + cases (exists_current … Hcur) #ls * #rs #Hintape #cmark + cases (is_marked_to_exists … cmark) #c #Hcm + >Hintape >Hcm -Hintape -Hcm #Hta + @(ex_intro … ls) @(ex_intro … c) @(ex_intro …rs) % [//] lapply Hta -Hta + (* #ls #c #rs #Hintape whd in Hta; + >Hintape in Hta; * #_ -Hintape forse non serve *) + cases (true_or_false (c==bit false)) #Hc + [>(\P Hc) #Hta % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 ?? (proj1 ?? (Htb … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (Htb … Hta) (refl …))) + ] + |cases (true_or_false (c==bit true)) #Hc1 + [>(\P Hc1) #Hta + cut (〈bit true, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta)(refl …))) + ] + |cases (true_or_false (c==null)) #Hc2 + [>(\P Hc2) #Hta + cut (〈null, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq + cut (〈null, true〉 ≠ 〈bit true, true〉) [% #Hdes destruct] #Hneq1 % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) + ] + |#Hta cut (bit_or_null c = false) + [lapply Hc; lapply Hc1; lapply Hc2 -Hc -Hc1 -Hc2 + cases c normalize [* normalize /2/] /2/] #Hcut % + [%[cases (Htb … Hta) #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc; destruct] #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc1; destruct] #_ -Htb #Htb + lapply (Htb ?) [% #H destruct (H) normalize in Hc2; destruct] + * #_ #Houttape #_ @(Houttape … Hta) + |>Hcut #H destruct + ] + |#l1 #c0 #l2 >Hcut #H destruct + ] + ] + ] + ] +|#intape #outape #ta #Hta #Htb #ls #c #rs #Hintape + >Hintape in Hta; whd in ⊢ (%→?); * #Hmark #Hta % [@Hmark //] + whd in Htb; >Htb // +] +qed. + +(* old universal version + +definition R_comp_step_true ≝ λt1,t2. + ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls 〈c,true〉 rs → + (* bit_or_null c = false *) + (bit_or_null c = false → t2 = midtape ? ls 〈c,false〉 rs) ∧ + (* no marks in rs *) + (bit_or_null c = true → + (∀c.memb ? c rs = true → is_marked ? c = false) → + ∀a,l. (a::l) = reverse ? (〈c,true〉::rs) → + t2 = rightof (FinProd FSUnialpha FinBool) a (l@ls)) ∧ + (∀l1,c0,l2. + bit_or_null c = true → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = l1@〈c0,true〉::l2 → + (c = c0 → + l2 = [ ] → (* test true but l2 is empty *) + t2 = rightof ? 〈c0,false〉 ((reverse ? l1)@〈c,true〉::ls)) ∧ + (c = c0 → + ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *) + 〈a,false〉::l1' = l1@[〈c0,false〉] → + l2 = 〈a0,b〉::l2' → + t2 = midtape ? (〈c,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧ + (c ≠ c0 →(* test false *) + t2 = midtape (FinProd … FSUnialpha FinBool) + ((reverse ? l1)@〈c,true〉::ls) 〈c0,false〉 l2)). + +definition R_comp_step_false ≝ + λt1,t2. + ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → + is_marked ? c = false ∧ t2 = t1. + +(* +lemma is_marked_to_exists: ∀alpha,c. is_marked alpha c = true → + ∃c'. c = 〈c',true〉. +#alpha * c *) + +lemma sem_comp_step : + accRealize ? comp_step (inr … (inl … (inr … start_nop))) + R_comp_step_true R_comp_step_false. +@(acc_sem_if_app … (sem_test_char ? (is_marked ?)) + (sem_comp_step_subcase FSUnialpha 〈bit false,true〉 ?? + (sem_comp_step_subcase FSUnialpha 〈bit true,true〉 ?? + (sem_comp_step_subcase FSUnialpha 〈null,true〉 ?? + (sem_clear_mark …)))) + (sem_nop …) …) +[#intape #outape #ta #Hta #Htb #ls #c #rs #Hintape whd in Hta; + >Hintape in Hta; * #_ -Hintape (* forse non serve *) + cases (true_or_false (c==bit false)) #Hc + [>(\P Hc) #Hta % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 ?? (proj1 ?? (Htb … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (Htb … Hta) (refl …))) + ] + |cases (true_or_false (c==bit true)) #Hc1 + [>(\P Hc1) #Hta + cut (〈bit true, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta)(refl …))) + ] + |cases (true_or_false (c==null)) #Hc2 + [>(\P Hc2) #Hta + cut (〈null, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq + cut (〈null, true〉 ≠ 〈bit true, true〉) [% #Hdes destruct] #Hneq1 % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) + ] + |#Hta cut (bit_or_null c = false) + [lapply Hc; lapply Hc1; lapply Hc2 -Hc -Hc1 -Hc2 + cases c normalize [* normalize /2/] /2/] #Hcut % + [%[cases (Htb … Hta) #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc; destruct] #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc1; destruct] #_ -Htb #Htb + lapply (Htb ?) [% #H destruct (H) normalize in Hc2; destruct] + * #_ #Houttape #_ @(Houttape … Hta) + |>Hcut #H destruct + ] + |#l1 #c0 #l2 >Hcut #H destruct + ] + ] + ] + ] +|#intape #outape #ta #Hta #Htb #ls #c #rs #Hintape + >Hintape in Hta; whd in ⊢ (%→?); * #Hmark #Hta % [@Hmark //] + whd in Htb; >Htb // +] +qed. *) + +(* definition R_comp_step_true ≝ λt1,t2. ∀l0,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) l0 c rs → @@ -704,7 +1258,7 @@ definition R_comp_step_false ≝ is_marked ? c = false ∧ t2 = t1. lemma sem_comp_step : - accRealize ? comp_step (inr … (inl … (inr … 0))) + accRealize ? comp_step (inr … (inl … (inr … start_nop))) R_comp_step_true R_comp_step_false. #intape cases (acc_sem_if … (sem_test_char ? (is_marked ?)) @@ -718,9 +1272,9 @@ cases (acc_sem_if … (sem_test_char ? (is_marked ?)) [ % [@Hloop ] ] -Hloop [ #Hstate lapply (H1 Hstate) -H1 -Hstate -H2 * #ta * whd in ⊢ (%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; #Hleft cases (Hleft ? (refl ??)) -Hleft - cases c in Hintape; #c' #b #Hintape whd in ⊢ (??%?→?); - #Hb >Hb #Hta @(ex_intro ?? c') % // + >Hintape in Hleft; * * + cases c in Hintape; #c' #b #Hintape #x * whd in ⊢ (??%?→?); #H destruct (H) + whd in ⊢ (??%?→?); #Hb >Hb #Hta @(ex_intro ?? c') % // cases (Hright … Hta) [ * #Hc' #H1 % % [destruct (Hc') % ] #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 @@ -762,14 +1316,12 @@ cases (acc_sem_if … (sem_test_char ? (is_marked ?)) ] | #Hstate lapply (H2 Hstate) -H1 -Hstate -H2 * #ta * whd in ⊢ (%→%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; #Hleft - cases (Hleft ? (refl ??)) -Hleft - #Hc #Hta % // >Hright // + >Hintape in Hleft; * #Hc #Hta % [@Hc % | >Hright //] ] -qed. +qed.*) definition compare ≝ - whileTM ? comp_step (inr … (inl … (inr … 0))). + whileTM ? comp_step (inr … (inl … (inr … start_nop))). (* definition R_compare := @@ -817,11 +1369,19 @@ RIFIUTO: c ≠ d t2 = midtape ? l0 〈grid,false〉 (l1@〈comma,true〉::l2)) ∨ (b = bit x ∧ b = c ∧ bs = b0s *) + +definition list_cases2: ∀A.∀P:list A →list A →Prop.∀l1,l2. |l1| = |l2| → +P [ ] [ ] → (∀a1,a2:A.∀tl1,tl2. |tl1| = |tl2| → P (a1::tl1) (a2::tl2)) → P l1 l2. +#A #P #l1 #l2 #Hlen lapply Hlen @(list_ind2 … Hlen) // +#tl1 #tl2 #hd1 #hd2 #Hind normalize #HlenS #H1 #H2 @H2 // +qed. + definition R_compare := λt1,t2. ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → (∀c'.bit_or_null c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧ (∀c'. c = 〈c',false〉 → t2 = t1) ∧ +(* forse manca il caso no marks in rs *) ∀b,b0,bs,b0s,l1,l2. |bs| = |b0s| → (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → bit_or_null (\fst c) = true) → @@ -845,32 +1405,6 @@ definition R_compare := reverse ? la@ls) 〈d',false〉 (lc@〈comma,false〉::l2)). -lemma list_ind2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop. - length ? l1 = length ? l2 → - (P [] []) → - (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → - P l1 l2. -#T1 #T2 #l1 #l2 #P #Hl #Pnil #Pcons -generalize in match Hl; generalize in match l2; -elim l1 -[#l2 cases l2 // normalize #t2 #tl2 #H destruct -|#t1 #tl1 #IH #l2 cases l2 - [normalize #H destruct - |#t2 #tl2 #H @Pcons @IH normalize in H; destruct // ] -] -qed. - -lemma list_cases_2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop. - length ? l1 = length ? l2 → - (l1 = [] → l2 = [] → P) → - (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P. -#T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl) -[ #Pnil #Pcons @Pnil // -| #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ] -qed. - lemma wsem_compare : WRealize ? compare R_compare. #t #i #outc #Hloop lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] @@ -885,28 +1419,85 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] cases (Rfalse … Htapea) -Rfalse >Hc whd in ⊢ (??%?→?);#Hfalse destruct (Hfalse) ] | #tapea #tapeb #tapec #Hleft #Hright #IH #Htapec lapply (IH Htapec) -Htapec -IH #IH - whd in Hleft; #ls #c #rs #Htapea cases (Hleft … Htapea) -Hleft - #c' * #Hc >Hc cases (true_or_false (bit_or_null c')) #Hc' - [2: * - [ * >Hc' #H @False_ind destruct (H) - | * #_ #Htapeb cases (IH … Htapeb) * #_ #H #_ % - [% - [#c1 #Hc1 #Heqc destruct (Heqc) Hc' #Hfalse @False_ind destruct (Hfalse) + whd in Hleft; #ls #c #rs #Htapea cases Hleft -Hleft + #ls0 * #c' * #rs0 * >Htapea #Hdes destruct (Hdes) * * + cases (true_or_false (bit_or_null c')) #Hc' + [2: #Htapeb lapply (Htapeb Hc') -Htapeb #Htapeb #_ #_ % + [%[#c1 #Hc1 #Heqc destruct (Heqc) + cases (IH … Htapeb) * #_ #H #_ Hc' #Hfalse @False_ind destruct (Hfalse) ] - |#Hleft % + |#_ (* no marks in rs ??? *) #_ #Hleft % [ % [ #c'' #Hc'' #Heq destruct (Heq) >Hc'' in Hc'; #H destruct (H) | #c0 #Hfalse destruct (Hfalse) ] |#b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1 - #Heq destruct (Heq) #_ #Hrs cases Hleft -Hleft + #Heq destruct (Heq) #_ >append_cons; (memb_single … memc1) %] + |@Hl1 @memc1 + ] + |* (* manca il caso in cui dopo una parte uguale il + secondo nastro finisca ... ???? *) + #_ cases (true_or_false (b==b0)) #eqbb0 + [2: #_ #Htapeb %2 lapply (Htapeb … (\Pf eqbb0)) -Htapeb #Htapeb + cases (IH … Htapeb) * #_ #Hout #_ + @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0) + @(ex_intro … bs) @(ex_intro … b0s) % + [%[%[@(\Pf eqbb0) | %] | %] + |>(Hout … (refl …)) -Hout >Htapeb @eq_f3 [2,3:%] + >reverse_append >reverse_append >associative_append + >associative_append % + ] + |lapply Hbs1 lapply Hbs2 lapply Hrs -Hbs1 -Hbs2 -Hrs + @(list_cases2 … Hlen) + [#Hrs #_ #_ >associative_append >associative_append #Htapeb #_ + lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb + cases (IH … Htapeb) -IH * #Hout #_ #_ %1 % + [>(\P eqbb0) % + |>(Hout grid (refl …) (refl …)) @eq_f + normalize >associative_append % + ] + |* #a1 #ba1 * #a2 #ba2 #tl1 #tl2 #HlenS #Hrs #Hbs1 #Hbs2 + cut (ba1 = false) [@(Hbs1 〈a1,ba1〉) @memb_hd] #Hba1 >Hba1 + >associative_append >associative_append #Htapeb #_ + lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb + cases (IH … Htapeb) -IH * #_ #_ + #IH cases(IH a1 a2 ?? (l1@[〈b0,false〉]) l2 HlenS ????? (refl …) ??) + [ + + +(* + cut (∃a,l1'.〈a,false〉::l1'=((bs@[〈grid,false〉])@l1)@[〈b0,false〉]) + [generalize in match Hbs2; cases bs + [#_ @(ex_intro … grid) @(ex_intro … (l1@[〈b0,false〉])) + >associative_append % + |* #bsc #bsb #bstl #Hbs2 @(ex_intro … bsc) + @(ex_intro … (((bstl@[〈grid,false〉])@l1)@[〈b0,false〉])) + normalize @eq_f2 [2:%] @eq_f @sym_eq @(Hbs2 〈bsc,bsb〉) @memb_hd + ] + ] + * #a * #l1' #H2 + cut (∃a0,b1,l2'.b0s@〈comma,false〉::l2=〈a0,b1〉::l2') + [cases b0s + [@(ex_intro … comma) @(ex_intro … false) @(ex_intro … l2) % + |* #bsc #bsb #bstl @(ex_intro … bsc) @(ex_intro … bsb) + @(ex_intro … (bstl@〈comma,false〉::l2)) % + ] + ] *) + * #a0 * #b1 * #l2' #H3 + lapply (Htapeb … (\P eqbb0) a a0 b1 l1' l2' H2 H3) -Htapeb #Htapeb + cases (IH … Htapeb) -IH * + + [2: * >Hc' #Hfalse @False_ind destruct ] * #_ - @(list_cases_2 … Hlen) + @(list_cases2 … Hlen) [ #Hbs #Hb0s generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?); -Hrs #Hrs normalize in Hrs; #Hleft cases (Hleft ????? Hrs ?) -Hleft [ * #Heqb #Htapeb cases (IH … Htapeb) -IH * #IH #_ #_ @@ -983,5 +1574,5 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] ] ]]]]] qed. - -axiom sem_compare : Realize ? compare R_compare. \ No newline at end of file + +axiom sem_compare : Realize ? compare R_compare.