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