# HG changeset patch # User Christian Urban # Date 1359516785 0 # Node ID fe7a257bdff4fe022d089e03a03cdb4293905146 # Parent 860f05037c36e704d8131f37d948f49bc56984dc updated paper diff -r 860f05037c36 -r fe7a257bdff4 Paper/Paper.thy --- a/Paper/Paper.thy Wed Jan 30 02:29:47 2013 +0000 +++ b/Paper/Paper.thy Wed Jan 30 03:33:05 2013 +0000 @@ -118,10 +118,10 @@ done lemmas HR1 = - Hoare_plus_halt_simple[where ?P1.0="P" and ?P2.0="Q" and ?P3.0="R\" and ?A="p\<^isub>1" and B="p\<^isub>2"] + Hoare_plus_halt[where ?S.0="R\" and ?A="p\<^isub>1" and B="p\<^isub>2"] lemmas HR2 = - Hoare_plus_unhalt_simple[where ?P1.0="P" and ?P2.0="Q" and ?A="p\<^isub>1" and B="p\<^isub>2"] + Hoare_plus_unhalt[where ?A="p\<^isub>1" and B="p\<^isub>2"] lemma inv_begin01: assumes "n > 1" @@ -765,6 +765,14 @@ \end{center} \noindent + For our Hoare-triples we can easily prove the following consequence rule + + \begin{equation} + @{thm[mode=Rule] Hoare_consequence} + \end{equation} + + + \noindent Like Asperti and Ricciotti with their notion of realisability, we have set up our Hoare-rules so that we can deal explicitly with total correctness and non-terminantion, rather than have @@ -928,7 +936,7 @@ \begin{center} - @{thm haltP_def[where lm="ns"]} + @{thm haltP_def} \end{center} diff -r 860f05037c36 -r fe7a257bdff4 paper.pdf Binary file paper.pdf has changed diff -r 860f05037c36 -r fe7a257bdff4 thys/turing_hoare.thy --- a/thys/turing_hoare.thy Wed Jan 30 02:29:47 2013 +0000 +++ b/thys/turing_hoare.thy Wed Jan 30 03:33:05 2013 +0000 @@ -10,6 +10,10 @@ where "P \ Q \ \l r. P (l, r) \ Q (l, r)" +lemma [intro, simp]: + "P \ P" +unfolding assert_imp_def by simp + fun holds_for :: "(tape \ bool) \ config \ bool" ("_ holds'_for _" [100, 99] 100) where @@ -57,24 +61,23 @@ text {* - {P1} A {Q1} {P2} B {Q2} Q1 \ P2 A well-formed - --------------------------------------------------- - {P1} A |+| B {Q2} + {P} A {Q} {Q} B {S} A well-formed + ----------------------------------------- + {P} A |+| B {S} *} -lemma Hoare_plus_halt [case_names A_halt B_halt Imp A_wf]: - assumes A_halt : "{P1} A {Q1}" - and B_halt : "{P2} B {Q2}" - and a_imp: "Q1 \ P2" +lemma Hoare_plus_halt [case_names A_halt B_halt A_wf]: + assumes A_halt : "{P} A {Q}" + and B_halt : "{Q} B {S}" and A_wf : "tm_wf (A, 0)" - shows "{P1} A |+| B {Q2}" + shows "{P} A |+| B {S}" proof(rule Hoare_haltI) fix l r - assume h: "P1 (l, r)" + assume h: "P (l, r)" then obtain n1 l' r' where "is_final (steps0 (1, l, r) A n1)" - and a1: "Q1 holds_for (steps0 (1, l, r) A n1)" + and a1: "Q holds_for (steps0 (1, l, r) A n1)" and a2: "steps0 (1, l, r) A n1 = (0, l', r')" using A_halt unfolding Hoare_halt_def by (metis is_final_eq surj_pair) @@ -82,48 +85,37 @@ where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')" using A_wf by (rule_tac tm_comp_pre_halt_same) moreover - from a1 a2 a_imp have "P2 (l', r')" by (simp add: assert_imp_def) + from a1 a2 have "Q (l', r')" by (simp) then obtain n3 l'' r'' where "is_final (steps0 (1, l', r') B n3)" - and b1: "Q2 holds_for (steps0 (1, l', r') B n3)" + and b1: "S holds_for (steps0 (1, l', r') B n3)" and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')" using B_halt unfolding Hoare_halt_def by (metis is_final_eq surj_pair) then have "steps0 (Suc (length A div 2), l', r') (A |+| B) n3 = (0, l'', r'')" using A_wf by (rule_tac tm_comp_second_halt_same) ultimately show - "\n. is_final (steps0 (1, l, r) (A |+| B) n) \ Q2 holds_for (steps0 (1, l, r) (A |+| B) n)" + "\n. is_final (steps0 (1, l, r) (A |+| B) n) \ S holds_for (steps0 (1, l, r) (A |+| B) n)" using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp) qed -lemma Hoare_plus_halt_simple [case_names A_halt B_halt A_wf]: - assumes A_halt : "{P1} A {P2}" - and B_halt : "{P2} B {P3}" - and A_wf : "tm_wf (A, 0)" - shows "{P1} A |+| B {P3}" -by (rule Hoare_plus_halt[OF A_halt B_halt _ A_wf]) - (simp add: assert_imp_def) - - - text {* - {P1} A {Q1} {P2} B loops Q1 \ P2 A well-formed - ------------------------------------------------------ - {P1} A |+| B loops + {P} A {Q} {Q} B loops A well-formed + ------------------------------------------ + {P} A |+| B loops *} -lemma Hoare_plus_unhalt [case_names A_halt B_unhalt Imp A_wf]: - assumes A_halt: "{P1} A {Q1}" - and B_uhalt: "{P2} B \" - and a_imp: "Q1 \ P2" +lemma Hoare_plus_unhalt [case_names A_halt B_unhalt A_wf]: + assumes A_halt: "{P} A {Q}" + and B_uhalt: "{Q} B \" and A_wf : "tm_wf (A, 0)" - shows "{P1} (A |+| B) \" + shows "{P} (A |+| B) \" proof(rule_tac Hoare_unhaltI) fix n l r - assume h: "P1 (l, r)" + assume h: "P (l, r)" then obtain n1 l' r' where a: "is_final (steps0 (1, l, r) A n1)" - and b: "Q1 holds_for (steps0 (1, l, r) A n1)" + and b: "Q holds_for (steps0 (1, l, r) A n1)" and c: "steps0 (1, l, r) A n1 = (0, l', r')" using A_halt unfolding Hoare_halt_def by (metis is_final_eq surj_pair) @@ -132,12 +124,12 @@ then show "\ is_final (steps0 (1, l, r) (A |+| B) n)" proof(cases "n2 \ n") case True - from b c a_imp have "P2 (l', r')" by (simp add: assert_imp_def) + from b c have "Q (l', r')" by simp then have "\ n. \ is_final (steps0 (1, l', r') B n) " using B_uhalt unfolding Hoare_unhalt_def by simp - then have "\ is_final (steps0 (Suc 0, l', r') B (n - n2))" by auto + then have "\ is_final (steps0 (1, l', r') B (n - n2))" by auto then obtain s'' l'' r'' - where "steps0 (Suc 0, l', r') B (n - n2) = (s'', l'', r'')" + where "steps0 (1, l', r') B (n - n2) = (s'', l'', r'')" and "\ is_final (s'', l'', r'')" by (metis surj_pair) then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')" using A_wf by (auto dest: tm_comp_second_same simp del: tm_wf.simps) @@ -155,19 +147,8 @@ qed qed -lemma Hoare_plus_unhalt_simple [case_names A_halt B_unhalt A_wf]: - assumes A_halt: "{P1} A {P2}" - and B_uhalt: "{P2} B \" - and A_wf : "tm_wf (A, 0)" - shows "{P1} (A |+| B) \" -by (rule Hoare_plus_unhalt[OF A_halt B_uhalt _ A_wf]) - (simp add: assert_imp_def) - - -lemma Hoare_weaken: - assumes a: "{P} p {Q}" - and b: "P' \ P" - and c: "Q \ Q'" +lemma Hoare_consequence: + assumes "P' \ P" "{P} p {Q}" "Q \ Q'" shows "{P'} p {Q'}" using assms unfolding Hoare_halt_def assert_imp_def diff -r 860f05037c36 -r fe7a257bdff4 thys/uncomputable.thy --- a/thys/uncomputable.thy Wed Jan 30 02:29:47 2013 +0000 +++ b/thys/uncomputable.thy Wed Jan 30 03:33:05 2013 +0000 @@ -895,16 +895,18 @@ proof - have "{inv_begin1 x} tcopy_begin {inv_begin0 x}" by (metis assms begin_correct) - moreover - have "{inv_loop1 x} tcopy_loop {inv_loop0 x}" - by (metis assms loop_correct) - moreover + moreover have "inv_begin0 x \ inv_loop1 x" unfolding assert_imp_def unfolding inv_begin0.simps inv_loop1.simps unfolding inv_loop1_loop.simps inv_loop1_exit.simps apply(auto simp add: numeral Cons_eq_append_conv) by (rule_tac x = "Suc 0" in exI, auto) + ultimately have "{inv_begin1 x} tcopy_begin {inv_loop1 x}" + by (rule_tac Hoare_consequence) (auto) + moreover + have "{inv_loop1 x} tcopy_loop {inv_loop0 x}" + by (metis assms loop_correct) ultimately have "{inv_begin1 x} (tcopy_begin |+| tcopy_loop) {inv_loop0 x}" by (rule_tac Hoare_plus_halt) (auto) @@ -912,7 +914,7 @@ have "{inv_end1 x} tcopy_end {inv_end0 x}" by (metis assms end_correct) moreover - have "inv_loop0 x \ inv_end1 x" + have "inv_loop0 x = inv_end1 x" by(auto simp: inv_end1.simps inv_loop1.simps assert_imp_def) ultimately show "{inv_begin1 x} tcopy {inv_end0 x}" @@ -955,10 +957,10 @@ (* invariants of dither *) abbreviation (input) - "dither_halt_inv \ \tp. (\n. tp = (Bk \ n, <1::nat>))" + "dither_halt_inv \ \tp. \k. tp = (Bk \ k, <1::nat>)" abbreviation (input) - "dither_unhalt_inv \ \tp. (\n. tp = (Bk \ n, <0::nat>))" + "dither_unhalt_inv \ \tp. \k. tp = (Bk \ k, <0::nat>)" lemma dither_loops_aux: "(steps0 (1, Bk \ m, [Oc]) dither stp = (1, Bk \ m, [Oc])) \ @@ -974,7 +976,6 @@ apply(auto simp add: numeral tape_of_nat_abv) by (metis Suc_neq_Zero is_final_eq) - lemma dither_halts_aux: shows "steps0 (1, Bk \ m, [Oc, Oc]) dither 2 = (0, Bk \ m, [Oc, Oc])" unfolding dither_def @@ -988,7 +989,6 @@ by (metis (lifting, mono_tags) holds_for.simps is_final_eq prod.cases) - section {* The diagnal argument below shows the undecidability of Halting problem *} text {* @@ -998,7 +998,7 @@ definition haltP :: "tprog0 \ nat list \ bool" where - "haltP p lm \ {(\tp. tp = ([], ))} p {(\tp. (\k n. tp = (Bk \ k, )))}" + "haltP p ns \ {(\tp. tp = ([], ))} p {(\tp. (\k n. tp = (Bk \ k, )))}" lemma [intro, simp]: "tm_wf0 tcopy" by (auto simp: tcopy_def) @@ -1026,9 +1026,9 @@ The following two assumptions specifies that @{text "H"} does solve the Halting problem. *} and h_case: - "\ M lm. haltP M lm \ {(\tp. tp = ([Bk], ))} H {(\tp. \n. tp = (Bk \ n, <0::nat>))}" + "\ M ns. haltP M ns \ {(\tp. tp = ([Bk], ))} H {(\tp. \k. tp = (Bk \ k, <0::nat>))}" and nh_case: - "\ M lm. \ haltP M lm \ {(\tp. tp = ([Bk], ))} H {(\tp. \n. tp = (Bk \ n, <1::nat>))}" + "\ M ns. \ haltP M ns \ {(\tp. tp = ([Bk], ))} H {(\tp. \k. tp = (Bk \ k, <1::nat>))}" begin (* invariants for H *) @@ -1036,10 +1036,10 @@ "pre_H_inv M n \ \tp. tp = ([Bk], <[code M, n]>)" abbreviation - "post_H_halt_inv \ \tp. \n. tp = (Bk \ n, <1::nat>)" + "post_H_halt_inv \ \tp. \k. tp = (Bk \ k, <1::nat>)" abbreviation - "post_H_unhalt_inv \ \tp. \n. tp = (Bk \ n, <0::nat>)" + "post_H_unhalt_inv \ \tp. \k. tp = (Bk \ k, <0::nat>)" lemma H_halt_inv: @@ -1067,7 +1067,7 @@ (* invariants *) def P1 \ "\tp. tp = ([]::cell list, <[code_tcontra]>)" def P2 \ "\tp. tp = ([Bk], <[code_tcontra, code_tcontra]>)" - def P3 \ "\tp. \n. tp = (Bk \ n, <1::nat>)" + def P3 \ "\tp. \k. tp = (Bk \ k, <1::nat>)" (* {P1} tcopy {P2} {P2} H {P3} @@ -1081,7 +1081,7 @@ (* {P1} (tcopy |+| H) {P3} *) have first: "{P1} (tcopy |+| H) {P3}" - proof (cases rule: Hoare_plus_halt_simple) + proof (cases rule: Hoare_plus_halt) case A_halt (* of tcopy *) show "{P1} tcopy {P2}" unfolding P1_def P2_def by (rule tcopy_correct) @@ -1099,7 +1099,7 @@ (* {P1} tcontra {P3} *) have "{P1} tcontra {P3}" unfolding tcontra_def - by (rule Hoare_plus_halt_simple[OF first second H_wf]) + by (rule Hoare_plus_halt[OF first second H_wf]) with assms show "False" unfolding P1_def P3_def @@ -1120,7 +1120,7 @@ (* invariants *) def P1 \ "\tp. tp = ([]::cell list, <[code_tcontra]>)" def P2 \ "\tp. tp = ([Bk], <[code_tcontra, code_tcontra]>)" - def Q3 \ "\tp. \n. tp = (Bk \ n, <0::nat>)" + def Q3 \ "\tp. \k. tp = (Bk \ k, <0::nat>)" (* {P1} tcopy {P2} {P2} H {Q3} @@ -1134,7 +1134,7 @@ (* {P1} (tcopy |+| H) {Q3} *) have first: "{P1} (tcopy |+| H) {Q3}" - proof (cases rule: Hoare_plus_halt_simple) + proof (cases rule: Hoare_plus_halt) case A_halt (* of tcopy *) show "{P1} tcopy {P2}" unfolding P1_def P2_def by (rule tcopy_correct) @@ -1152,7 +1152,7 @@ (* {P1} tcontra loops *) have "{P1} tcontra \" unfolding tcontra_def - by (rule Hoare_plus_unhalt_simple[OF first second H_wf]) + by (rule Hoare_plus_unhalt[OF first second H_wf]) with assms show "False" unfolding P1_def