55 |
59 |
56 |
60 |
57 |
61 |
58 |
62 |
59 text {* |
63 text {* |
60 {P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2 A well-formed |
64 {P} A {Q} {Q} B {S} A well-formed |
61 --------------------------------------------------- |
65 ----------------------------------------- |
62 {P1} A |+| B {Q2} |
66 {P} A |+| B {S} |
63 *} |
67 *} |
64 |
68 |
65 |
69 |
66 lemma Hoare_plus_halt [case_names A_halt B_halt Imp A_wf]: |
70 lemma Hoare_plus_halt [case_names A_halt B_halt A_wf]: |
67 assumes A_halt : "{P1} A {Q1}" |
71 assumes A_halt : "{P} A {Q}" |
68 and B_halt : "{P2} B {Q2}" |
72 and B_halt : "{Q} B {S}" |
69 and a_imp: "Q1 \<mapsto> P2" |
|
70 and A_wf : "tm_wf (A, 0)" |
73 and A_wf : "tm_wf (A, 0)" |
71 shows "{P1} A |+| B {Q2}" |
74 shows "{P} A |+| B {S}" |
72 proof(rule Hoare_haltI) |
75 proof(rule Hoare_haltI) |
73 fix l r |
76 fix l r |
74 assume h: "P1 (l, r)" |
77 assume h: "P (l, r)" |
75 then obtain n1 l' r' |
78 then obtain n1 l' r' |
76 where "is_final (steps0 (1, l, r) A n1)" |
79 where "is_final (steps0 (1, l, r) A n1)" |
77 and a1: "Q1 holds_for (steps0 (1, l, r) A n1)" |
80 and a1: "Q holds_for (steps0 (1, l, r) A n1)" |
78 and a2: "steps0 (1, l, r) A n1 = (0, l', r')" |
81 and a2: "steps0 (1, l, r) A n1 = (0, l', r')" |
79 using A_halt unfolding Hoare_halt_def |
82 using A_halt unfolding Hoare_halt_def |
80 by (metis is_final_eq surj_pair) |
83 by (metis is_final_eq surj_pair) |
81 then obtain n2 |
84 then obtain n2 |
82 where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')" |
85 where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')" |
83 using A_wf by (rule_tac tm_comp_pre_halt_same) |
86 using A_wf by (rule_tac tm_comp_pre_halt_same) |
84 moreover |
87 moreover |
85 from a1 a2 a_imp have "P2 (l', r')" by (simp add: assert_imp_def) |
88 from a1 a2 have "Q (l', r')" by (simp) |
86 then obtain n3 l'' r'' |
89 then obtain n3 l'' r'' |
87 where "is_final (steps0 (1, l', r') B n3)" |
90 where "is_final (steps0 (1, l', r') B n3)" |
88 and b1: "Q2 holds_for (steps0 (1, l', r') B n3)" |
91 and b1: "S holds_for (steps0 (1, l', r') B n3)" |
89 and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')" |
92 and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')" |
90 using B_halt unfolding Hoare_halt_def |
93 using B_halt unfolding Hoare_halt_def |
91 by (metis is_final_eq surj_pair) |
94 by (metis is_final_eq surj_pair) |
92 then have "steps0 (Suc (length A div 2), l', r') (A |+| B) n3 = (0, l'', r'')" |
95 then have "steps0 (Suc (length A div 2), l', r') (A |+| B) n3 = (0, l'', r'')" |
93 using A_wf by (rule_tac tm_comp_second_halt_same) |
96 using A_wf by (rule_tac tm_comp_second_halt_same) |
94 ultimately show |
97 ultimately show |
95 "\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps0 (1, l, r) (A |+| B) n)" |
98 "\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> S holds_for (steps0 (1, l, r) (A |+| B) n)" |
96 using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp) |
99 using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp) |
97 qed |
100 qed |
98 |
101 |
99 lemma Hoare_plus_halt_simple [case_names A_halt B_halt A_wf]: |
|
100 assumes A_halt : "{P1} A {P2}" |
|
101 and B_halt : "{P2} B {P3}" |
|
102 and A_wf : "tm_wf (A, 0)" |
|
103 shows "{P1} A |+| B {P3}" |
|
104 by (rule Hoare_plus_halt[OF A_halt B_halt _ A_wf]) |
|
105 (simp add: assert_imp_def) |
|
106 |
|
107 |
|
108 |
|
109 text {* |
102 text {* |
110 {P1} A {Q1} {P2} B loops Q1 \<mapsto> P2 A well-formed |
103 {P} A {Q} {Q} B loops A well-formed |
111 ------------------------------------------------------ |
104 ------------------------------------------ |
112 {P1} A |+| B loops |
105 {P} A |+| B loops |
113 *} |
106 *} |
114 |
107 |
115 lemma Hoare_plus_unhalt [case_names A_halt B_unhalt Imp A_wf]: |
108 lemma Hoare_plus_unhalt [case_names A_halt B_unhalt A_wf]: |
116 assumes A_halt: "{P1} A {Q1}" |
109 assumes A_halt: "{P} A {Q}" |
117 and B_uhalt: "{P2} B \<up>" |
110 and B_uhalt: "{Q} B \<up>" |
118 and a_imp: "Q1 \<mapsto> P2" |
|
119 and A_wf : "tm_wf (A, 0)" |
111 and A_wf : "tm_wf (A, 0)" |
120 shows "{P1} (A |+| B) \<up>" |
112 shows "{P} (A |+| B) \<up>" |
121 proof(rule_tac Hoare_unhaltI) |
113 proof(rule_tac Hoare_unhaltI) |
122 fix n l r |
114 fix n l r |
123 assume h: "P1 (l, r)" |
115 assume h: "P (l, r)" |
124 then obtain n1 l' r' |
116 then obtain n1 l' r' |
125 where a: "is_final (steps0 (1, l, r) A n1)" |
117 where a: "is_final (steps0 (1, l, r) A n1)" |
126 and b: "Q1 holds_for (steps0 (1, l, r) A n1)" |
118 and b: "Q holds_for (steps0 (1, l, r) A n1)" |
127 and c: "steps0 (1, l, r) A n1 = (0, l', r')" |
119 and c: "steps0 (1, l, r) A n1 = (0, l', r')" |
128 using A_halt unfolding Hoare_halt_def |
120 using A_halt unfolding Hoare_halt_def |
129 by (metis is_final_eq surj_pair) |
121 by (metis is_final_eq surj_pair) |
130 then obtain n2 where eq: "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')" |
122 then obtain n2 where eq: "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')" |
131 using A_wf by (rule_tac tm_comp_pre_halt_same) |
123 using A_wf by (rule_tac tm_comp_pre_halt_same) |
132 then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
124 then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
133 proof(cases "n2 \<le> n") |
125 proof(cases "n2 \<le> n") |
134 case True |
126 case True |
135 from b c a_imp have "P2 (l', r')" by (simp add: assert_imp_def) |
127 from b c have "Q (l', r')" by simp |
136 then have "\<forall> n. \<not> is_final (steps0 (1, l', r') B n) " |
128 then have "\<forall> n. \<not> is_final (steps0 (1, l', r') B n) " |
137 using B_uhalt unfolding Hoare_unhalt_def by simp |
129 using B_uhalt unfolding Hoare_unhalt_def by simp |
138 then have "\<not> is_final (steps0 (Suc 0, l', r') B (n - n2))" by auto |
130 then have "\<not> is_final (steps0 (1, l', r') B (n - n2))" by auto |
139 then obtain s'' l'' r'' |
131 then obtain s'' l'' r'' |
140 where "steps0 (Suc 0, l', r') B (n - n2) = (s'', l'', r'')" |
132 where "steps0 (1, l', r') B (n - n2) = (s'', l'', r'')" |
141 and "\<not> is_final (s'', l'', r'')" by (metis surj_pair) |
133 and "\<not> is_final (s'', l'', r'')" by (metis surj_pair) |
142 then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')" |
134 then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')" |
143 using A_wf by (auto dest: tm_comp_second_same simp del: tm_wf.simps) |
135 using A_wf by (auto dest: tm_comp_second_same simp del: tm_wf.simps) |
144 then have "\<not> is_final (steps0 (1, l, r) (A |+| B) (n2 + (n - n2)))" |
136 then have "\<not> is_final (steps0 (1, l, r) (A |+| B) (n2 + (n - n2)))" |
145 using A_wf by (simp only: steps_add eq) (simp add: tm_wf.simps) |
137 using A_wf by (simp only: steps_add eq) (simp add: tm_wf.simps) |