diff -r 01f688735b9b -r 2cae8a39803e thys/uncomputable.thy --- a/thys/uncomputable.thy Thu Jan 31 11:35:16 2013 +0000 +++ b/thys/uncomputable.thy Fri Feb 01 01:34:37 2013 +0000 @@ -63,8 +63,6 @@ | "inv_begin3 n (l, r) = (n > 0 \ (l, tl r) = (Bk # Oc \ n, []))" | "inv_begin4 n (l, r) = (n > 0 \ (l, r) = (Oc \ n, [Bk, Oc]) \ (l, r) = (Oc \ (n - 1), [Oc, Bk, Oc]))" - - fun inv_begin :: "nat \ config \ bool" where "inv_begin n (s, tp) = @@ -80,9 +78,9 @@ by (rule_tac x = "Suc i" in exI, simp) lemma inv_begin_step: - assumes "inv_begin x cf" - and "x > 0" - shows "inv_begin x (step0 cf tcopy_begin)" + assumes "inv_begin n cf" + and "n > 0" + shows "inv_begin n (step0 cf tcopy_begin)" using assms unfolding tcopy_begin_def apply(case_tac cf) @@ -94,9 +92,9 @@ done lemma inv_begin_steps: - assumes "inv_begin x cf" - and "x > 0" - shows "inv_begin x (steps0 cf tcopy_begin stp)" + assumes "inv_begin n cf" + and "n > 0" + shows "inv_begin n (steps0 cf tcopy_begin stp)" apply(induct stp) apply(simp add: assms) apply(auto simp del: steps.simps) @@ -104,23 +102,22 @@ apply(simp_all add: assms) done -(* lemma begin_partial_correctness: - assumes "\stp. is_final (steps0 (1, [], Oc \ n) tcopy_begin stp)" + assumes "is_final (steps0 (1, [], Oc \ n) tcopy_begin stp)" shows "0 < n \ {inv_begin1 n} tcopy_begin {inv_begin0 n}" -using assms -apply(auto simp add: Hoare_halt_def) -apply(rule_tac x="stp" in exI) -apply(simp) -apply(case_tac "steps0 (Suc 0, [], Oc \ n) tcopy_begin stp") -apply(simp) -apply(case_tac n) -apply(simp) -apply(simp) -apply(case_tac nat) -apply(simp) -apply(simp) -*) +proof(rule_tac Hoare_haltI) + fix l r + assume h: "0 < n" "inv_begin1 n (l, r)" + have "inv_begin n (steps0 (1, [], Oc \ n) tcopy_begin stp)" + using h by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps) + then show + "\stp. is_final (steps0 (1, l, r) tcopy_begin stp) \ + inv_begin0 n holds_for steps (1, l, r) (tcopy_begin, 0) stp" + using h assms + apply(rule_tac x = stp in exI) + apply(case_tac "(steps0 (1, [], Oc \ n) tcopy_begin stp)", simp add: inv_begin.simps) + done +qed fun measure_begin_state :: "config \ nat" where @@ -135,69 +132,44 @@ else 0)" definition - "begin_LE = measures [measure_begin_state, measure_begin_step]" + "measure_begin = measures [measure_begin_state, measure_begin_step]" -lemma [simp]: "\tl r = []; r \ []; r \ [Bk]\ \ r = [Oc]" -by (case_tac r, auto, case_tac a, auto) +lemma wf_measure_begin: + shows "wf measure_begin" +unfolding measure_begin_def +by auto - -lemma halt_lemma: - "\wf LE; \n. (\ P (f n) \ (f (Suc n), (f n)) \ LE)\ \ \n. P (f n)" +lemma measure_begin_induct [case_names Step]: + "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_begin\ \ \n. P (f n)" +using wf_measure_begin by (metis wf_iff_no_infinite_down_chain) -lemma begin_halt: - "x > 0 \ \ stp. is_final (steps0 (1, [], Oc \ x) tcopy_begin stp)" -proof(rule_tac LE = begin_LE in halt_lemma) - show "wf begin_LE" unfolding begin_LE_def by (auto) -next - assume h: "0 < x" - show "\n. \ is_final (steps (1, [], Oc \ x) (tcopy_begin, 0) n) \ - (steps (1, [], Oc \ x) (tcopy_begin, 0) (Suc n), - steps (1, [], Oc \ x) (tcopy_begin, 0) n) \ begin_LE" - proof(rule_tac allI, rule_tac impI) - fix n - assume a: "\ is_final (steps (1, [], Oc \ x) (tcopy_begin, 0) n)" - have b: "inv_begin x (steps (1, [], Oc \ x) (tcopy_begin, 0) n)" - apply(rule_tac inv_begin_steps) - apply(simp_all add: inv_begin.simps h) - done - obtain s l r where c: "(steps (1, [], Oc \ x) (tcopy_begin, 0) n) = (s, l, r)" - apply(case_tac "steps (1, [], Oc \ x) (tcopy_begin, 0) n", auto) - done - moreover hence "inv_begin x (s, l, r)" - using c b by simp - ultimately show "(steps (1, [], Oc \ x) (tcopy_begin, 0) (Suc n), - steps (1, [], Oc \ x) (tcopy_begin, 0) n) \ begin_LE" - using a - proof(simp del: inv_begin.simps step.simps steps.simps) - assume "inv_begin x (s, l, r)" "0 < s" - thus "(step (s, l, r) (tcopy_begin, 0), s, l, r) \ begin_LE" - apply(auto simp: begin_LE_def step.simps tcopy_begin_def numeral split: if_splits) - apply(case_tac r, auto, case_tac a, auto) - done - qed - qed +lemma begin_halts: + assumes h: "x > 0" + shows "\ stp. is_final (steps0 (1, [], Oc \ x) tcopy_begin stp)" +proof (induct rule: measure_begin_induct) + case (Step n) + have "\ is_final (steps0 (1, [], Oc \ x) tcopy_begin n)" by fact + moreover + have "inv_begin x (steps0 (1, [], Oc \ x) tcopy_begin n)" + by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps h) + moreover + obtain s l r where eq: "(steps0 (1, [], Oc \ x) tcopy_begin n) = (s, l, r)" + by (metis measure_begin_state.cases) + ultimately + have "(step0 (s, l, r) tcopy_begin, s, l, r) \ measure_begin" + apply(auto simp: measure_begin_def tcopy_begin_def numeral split: if_splits) + apply(subgoal_tac "r = [Oc]") + apply(auto) + by (metis cell.exhaust list.exhaust tl.simps(2)) + then + show "(steps0 (1, [], Oc \ x) tcopy_begin (Suc n), steps0 (1, [], Oc \ x) tcopy_begin n) \ measure_begin" + using eq by (simp only: step_red) qed - + lemma begin_correct: - "0 < n \ {inv_begin1 n} tcopy_begin {inv_begin0 n}" -proof(rule_tac Hoare_haltI) - fix l r - assume h: "0 < n" "inv_begin1 n (l, r)" - then have "\ stp. is_final (steps0 (1, [], Oc \ n) tcopy_begin stp)" - by (rule_tac begin_halt) - then obtain stp where "is_final (steps0 (1, [], Oc \ n) tcopy_begin stp)" .. - moreover have "inv_begin n (steps0 (1, [], Oc \ n) tcopy_begin stp)" - apply(rule_tac inv_begin_steps) - using h by(simp_all add: inv_begin.simps) - ultimately show - "\stp. is_final (steps0 (1, l, r) tcopy_begin stp) \ - inv_begin0 n holds_for steps (1, l, r) (tcopy_begin, 0) stp" - using h - apply(rule_tac x = stp in exI) - apply(case_tac "(steps0 (1, [], Oc \ n) tcopy_begin stp)", simp add: inv_begin.simps) - done -qed + shows "0 < n \ {inv_begin1 n} tcopy_begin {inv_begin0 n}" +using begin_partial_correctness begin_halts by blast declare tm_comp.simps [simp del] declare adjust.simps[simp del] @@ -206,7 +178,6 @@ declare step.simps[simp del] declare steps.simps[simp del] - (* tcopy_loop *) fun @@ -217,8 +188,8 @@ inv_loop6_loop :: "nat \ tape \ bool" and inv_loop6_exit :: "nat \ tape \ bool" where - "inv_loop1_loop x (l, r) = (\ i j. i + j + 1 = x \ (l, r) = (Oc\i, Oc#Oc#Bk\j @ Oc\j) \ j > 0)" -| "inv_loop1_exit x (l, r) = ((l, r) = ([], Bk#Oc#Bk\x @ Oc\x) \ x > 0)" + "inv_loop1_loop n (l, r) = (\ i j. i + j + 1 = n \ (l, r) = (Oc\i, Oc#Oc#Bk\j @ Oc\j) \ j > 0)" +| "inv_loop1_exit n (l, r) = (0 < n \ (l, r) = ([], Bk#Oc#Bk\n @ Oc\n))" | "inv_loop5_loop x (l, r) = (\ i j k t. i + j = Suc x \ i > 0 \ j > 0 \ k + t = j \ t > 0 \ (l, r) = (Oc\k@Bk\j@Oc\i, Oc\t))" | "inv_loop5_exit x (l, r) = @@ -237,15 +208,15 @@ inv_loop5 :: "nat \ tape \ bool" and inv_loop6 :: "nat \ tape \ bool" where - "inv_loop0 x (l, r) = ((l, r) = ([Bk], Oc # Bk\x @ Oc\x) \ x > 0)" -| "inv_loop1 x (l, r) = (inv_loop1_loop x (l, r) \ inv_loop1_exit x (l, r))" -| "inv_loop2 x (l, r) = (\ i j any. i + j = x \ x > 0 \ i > 0 \ j > 0 \ (l, r) = (Oc\i, any#Bk\j@Oc\j))" -| "inv_loop3 x (l, r) = - (\ i j k t. i + j = x \ i > 0 \ j > 0 \ k + t = Suc j \ (l, r) = (Bk\k@Oc\i, Bk\t@Oc\j))" -| "inv_loop4 x (l, r) = - (\ i j k t. i + j = x \ i > 0 \ j > 0 \ k + t = j \ (l, r) = (Oc\k @ Bk\(Suc j)@Oc\i, Oc\t))" -| "inv_loop5 x (l, r) = (inv_loop5_loop x (l, r) \ inv_loop5_exit x (l, r))" -| "inv_loop6 x (l, r) = (inv_loop6_loop x (l, r) \ inv_loop6_exit x (l, r))" + "inv_loop0 n (l, r) = (0 < n \ (l, r) = ([Bk], Oc # Bk\n @ Oc\n))" +| "inv_loop1 n (l, r) = (inv_loop1_loop n (l, r) \ inv_loop1_exit n (l, r))" +| "inv_loop2 n (l, r) = (\ i j any. i + j = n \ n > 0 \ i > 0 \ j > 0 \ (l, r) = (Oc\i, any#Bk\j@Oc\j))" +| "inv_loop3 n (l, r) = + (\ i j k t. i + j = n \ i > 0 \ j > 0 \ k + t = Suc j \ (l, r) = (Bk\k@Oc\i, Bk\t@Oc\j))" +| "inv_loop4 n (l, r) = + (\ i j k t. i + j = n \ i > 0 \ j > 0 \ k + t = j \ (l, r) = (Oc\k @ Bk\(Suc j)@Oc\i, Oc\t))" +| "inv_loop5 n (l, r) = (inv_loop5_loop n (l, r) \ inv_loop5_exit n (l, r))" +| "inv_loop6 n (l, r) = (inv_loop6_loop n (l, r) \ inv_loop6_exit n (l, r))" fun inv_loop :: "nat \ config \ bool" where @@ -464,8 +435,7 @@ done lemma inv_loop_steps: - "\inv_loop x cf; x > 0\ - \ inv_loop x (steps cf (tcopy_loop, 0) stp)" + "\inv_loop x cf; x > 0\ \ inv_loop x (steps cf (tcopy_loop, 0) stp)" apply(induct stp, simp add: steps.simps, simp) apply(erule_tac inv_loop_step, simp) done @@ -489,14 +459,21 @@ else if s = 6 then length l else 0)" -definition loop_LE :: "(config \ config) set" +definition measure_loop :: "(config \ config) set" where - "loop_LE = measures [loop_stage, loop_state, loop_step]" + "measure_loop = measures [loop_stage, loop_state, loop_step]" -lemma wf_loop_le: "wf loop_LE" -unfolding loop_LE_def +lemma wf_measure_loop: "wf measure_loop" +unfolding measure_loop_def by (auto) +lemma measure_loop_induct [case_names Step]: + "\\n. \ P (f n) \ (f (Suc n), (f n)) \ measure_loop\ \ \n. P (f n)" +using wf_measure_loop +by (metis wf_iff_no_infinite_down_chain) + + + lemma [simp]: "inv_loop2 x ([], b) = False" by (auto simp: inv_loop2.simps) @@ -509,6 +486,7 @@ lemma [simp]: "inv_loop4 x ([], b) = False" by (auto simp: inv_loop4.simps) + lemma [elim]: "\inv_loop4 x (l', []); l' \ []; x > 0; length (takeWhile (\a. a = Oc) (rev l' @ [Oc])) \ @@ -581,7 +559,6 @@ apply(case_tac l', auto) done - lemma[elim]: "\inv_loop6 x (l', Oc # list); l' \ []; 0 < x; length (takeWhile (\a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) @@ -592,63 +569,58 @@ apply(auto simp: inv_loop6.simps) done -lemma loop_halt: - "\x > 0; inv_loop x (Suc 0, l, r)\ \ - \ stp. is_final (steps (Suc 0, l, r) (tcopy_loop, 0) stp)" -proof(rule_tac LE = loop_LE in halt_lemma) - show "wf loop_LE" by(intro wf_loop_le) -next - assume h: "0 < x" and g: "inv_loop x (Suc 0, l, r)" - show "\n. \ is_final (steps (Suc 0, l, r) (tcopy_loop, 0) n) \ - (steps (Suc 0, l, r) (tcopy_loop, 0) (Suc n), steps (Suc 0, l, r) (tcopy_loop, 0) n) \ loop_LE" - proof(rule_tac allI, rule_tac impI) - fix n - assume a: "\ is_final (steps (Suc 0, l, r) (tcopy_loop, 0) n)" - obtain s' l' r' where d: "(steps (Suc 0, l, r) (tcopy_loop, 0) n) = (s', l', r')" - by(case_tac "(steps (Suc 0, l, r) (tcopy_loop, 0) n)", auto) - hence "inv_loop x (s', l', r') \ s' \ 0" - using h g - apply(drule_tac stp = n in inv_loop_steps, auto) - using a - apply(simp) - done - hence "(step(s', l', r') (tcopy_loop, 0), s', l', r') \ loop_LE" - using h - apply(case_tac r', case_tac [2] a) - apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral loop_LE_def split: if_splits) - done - thus "(steps (Suc 0, l, r) (tcopy_loop, 0) (Suc n), - steps (Suc 0, l, r) (tcopy_loop, 0) n) \ loop_LE" - using d - apply(simp add: step_red) - done - qed +lemma loop_halts: + assumes h: "n > 0" "inv_loop n (1, l, r)" + shows "\ stp. is_final (steps0 (1, l, r) tcopy_loop stp)" +proof (induct rule: measure_loop_induct) + case (Step stp) + have "\ is_final (steps0 (1, l, r) tcopy_loop stp)" by fact + moreover + have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)" + by (rule_tac inv_loop_steps) (simp_all only: h) + moreover + obtain s l' r' where eq: "(steps0 (1, l, r) tcopy_loop stp) = (s, l', r')" + by (metis measure_begin_state.cases) + ultimately + have "(step0 (s, l', r') tcopy_loop, s, l', r') \ measure_loop" + using h(1) + apply(case_tac r') + apply(case_tac [2] a) + apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral measure_loop_def split: if_splits) + done + then + show "(steps0 (1, l, r) tcopy_loop (Suc stp), steps0 (1, l, r) tcopy_loop stp) \ measure_loop" + using eq by (simp only: step_red) qed lemma loop_correct: - "x > 0 \ {inv_loop1 x} tcopy_loop {inv_loop0 x}" + shows "0 < n \ {inv_loop1 n} tcopy_loop {inv_loop0 n}" + using assms proof(rule_tac Hoare_haltI) fix l r - assume h: "0 < x" - "inv_loop1 x (l, r)" - hence "\ stp. is_final (steps (Suc 0, l, r) (tcopy_loop, 0) stp)" - apply(rule_tac loop_halt, simp_all add: inv_loop.simps) + assume h: "0 < n" "inv_loop1 n (l, r)" + then obtain stp where k: "is_final (steps0 (1, l, r) tcopy_loop stp)" + using loop_halts + apply(simp add: inv_loop.simps) + apply(blast) done - then obtain stp where "is_final (steps (Suc 0, l, r) (tcopy_loop, 0) stp)" .. - moreover have "inv_loop x (steps (Suc 0, l, r) (tcopy_loop, 0) stp)" - apply(rule_tac inv_loop_steps) - using h by(simp_all add: inv_loop.simps) + moreover + have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)" + using h + by (rule_tac inv_loop_steps) (simp_all add: inv_loop.simps) ultimately show - "\n. is_final (steps (1, l, r) (tcopy_loop, 0) n) \ - inv_loop0 x holds_for steps (1, l, r) (tcopy_loop, 0) n" - using h + "\stp. is_final (steps0 (1, l, r) tcopy_loop stp) \ + inv_loop0 n holds_for steps0 (1, l, r) tcopy_loop stp" + using h(1) apply(rule_tac x = stp in exI) - apply(case_tac "(steps (Suc 0, l, r) (tcopy_loop, 0) stp)", - simp add: inv_begin.simps inv_loop.simps) + apply(case_tac "(steps0 (1, l, r) tcopy_loop stp)") + apply(simp add: inv_loop.simps) done qed + + (* tcopy_end *) fun @@ -664,27 +636,26 @@ inv_end1 :: "nat \ tape \ bool" and inv_end2 :: "nat \ tape \ bool" and inv_end3 :: "nat \ tape \ bool" and - inv_end4 :: "nat \ tape \ bool" and - + inv_end4 :: "nat \ tape \ bool" and inv_end5 :: "nat \ tape \ bool" where - "inv_end0 x (l, r) = (x > 0 \ l = [Bk] \ r = Oc\x @ Bk # Oc\x)" -| "inv_end1 x (l, r) = (x > 0 \ l = [Bk] \ r = Oc # Bk\x @ Oc\x)" -| "inv_end2 x (l, r) = (\ i j. i + j = Suc x \ x > 0 \ l = Oc\i @ [Bk] \ r = Bk\j @ Oc\x)" -| "inv_end3 x (l, r) = - (\ i j. x > 0 \ i + j = x \ l = Oc\i @ [Bk] \ r = Oc # Bk\j@ Oc\x)" -| "inv_end4 x (l, r) = (\ any. x > 0 \ l = Oc\x @ [Bk] \ r = any#Oc\x)" -| "inv_end5 x (l, r) = (inv_end5_loop x (l, r) \ inv_end5_exit x (l, r))" + "inv_end0 n (l, r) = (n > 0 \ (l, r) = ([Bk], Oc\n @ Bk # Oc\n))" +| "inv_end1 n (l, r) = (n > 0 \ (l, r) = ([Bk], Oc # Bk\n @ Oc\n))" +| "inv_end2 n (l, r) = (\ i j. i + j = Suc n \ n > 0 \ l = Oc\i @ [Bk] \ r = Bk\j @ Oc\n)" +| "inv_end3 n (l, r) = + (\ i j. n > 0 \ i + j = n \ l = Oc\i @ [Bk] \ r = Oc # Bk\j@ Oc\n)" +| "inv_end4 n (l, r) = (\ any. n > 0 \ l = Oc\n @ [Bk] \ r = any#Oc\n)" +| "inv_end5 n (l, r) = (inv_end5_loop n (l, r) \ inv_end5_exit n (l, r))" fun inv_end :: "nat \ config \ bool" where - "inv_end x (s, l, r) = (if s = 0 then inv_end0 x (l, r) - else if s = 1 then inv_end1 x (l, r) - else if s = 2 then inv_end2 x (l, r) - else if s = 3 then inv_end3 x (l, r) - else if s = 4 then inv_end4 x (l, r) - else if s = 5 then inv_end5 x (l, r) + "inv_end n (s, l, r) = (if s = 0 then inv_end0 n (l, r) + else if s = 1 then inv_end1 n (l, r) + else if s = 2 then inv_end2 n (l, r) + else if s = 3 then inv_end3 n (l, r) + else if s = 4 then inv_end4 n (l, r) + else if s = 5 then inv_end5 n (l, r) else False)" declare inv_end.simps[simp del] inv_end1.simps[simp del] @@ -843,6 +814,10 @@ lemma [simp]: "inv_end5 x ([], Oc # list) = False" by (auto simp: inv_end5.simps inv_end5_loop.simps) +lemma halt_lemma: + "\wf LE; \n. (\ P (f n) \ (f (Suc n), (f n)) \ LE)\ \ \n. P (f n)" +by (metis wf_iff_no_infinite_down_chain) + lemma end_halt: "\x > 0; inv_end x (Suc 0, l, r)\ \ \ stp. is_final (steps (Suc 0, l, r) (tcopy_end, 0) stp)" @@ -875,20 +850,20 @@ qed lemma end_correct: - "x > 0 \ {inv_end1 x} tcopy_end {inv_end0 x}" + "n > 0 \ {inv_end1 n} tcopy_end {inv_end0 n}" proof(rule_tac Hoare_haltI) fix l r - assume h: "0 < x" - "inv_end1 x (l, r)" + assume h: "0 < n" + "inv_end1 n (l, r)" then have "\ stp. is_final (steps0 (1, l, r) tcopy_end stp)" by (simp add: end_halt inv_end.simps) then obtain stp where "is_final (steps0 (1, l, r) tcopy_end stp)" .. - moreover have "inv_end x (steps0 (1, l, r) tcopy_end stp)" + moreover have "inv_end n (steps0 (1, l, r) tcopy_end stp)" apply(rule_tac inv_end_steps) using h by(simp_all add: inv_end.simps) ultimately show - "\n. is_final (steps (1, l, r) (tcopy_end, 0) n) \ - inv_end0 x holds_for steps (1, l, r) (tcopy_end, 0) n" + "\stp. is_final (steps (1, l, r) (tcopy_end, 0) stp) \ + inv_end0 n holds_for steps (1, l, r) (tcopy_end, 0) stp" using h apply(rule_tac x = stp in exI) apply(case_tac "(steps0 (1, l, r) tcopy_end stp)") @@ -940,10 +915,10 @@ by (rule_tac Hoare_plus_halt) (auto) qed -abbreviation - "pre_tcopy n \ \tp. tp = ([]::cell list, <[n::nat]>)" -abbreviation - "post_tcopy n \ \tp. tp= ([Bk], <[n, n::nat]>)" +abbreviation (input) + "pre_tcopy n \ \tp. tp = ([]::cell list, <(n::nat)>)" +abbreviation (input) + "post_tcopy n \ \tp. tp= ([Bk], <(n, n::nat)>)" lemma tcopy_correct: shows "{pre_tcopy n} tcopy {post_tcopy n}" @@ -955,7 +930,7 @@ by (auto simp add: tape_of_nl_abv tape_of_nat_abv) moreover have "inv_end0 (Suc n) = post_tcopy n" - by (auto simp add: inv_end0.simps tape_of_nat_abv tape_of_nl_abv) + by (auto simp add: inv_end0.simps tape_of_nat_abv tape_of_nat_pair) ultimately show "{pre_tcopy n} tcopy {post_tcopy n}" by simp @@ -1044,30 +1019,30 @@ The following two assumptions specifies that @{text "H"} does solve the Halting problem. *} and h_case: - "\ M ns. haltP M ns \ {(\tp. tp = ([Bk], ))} H {(\tp. \k. tp = (Bk \ k, <0::nat>))}" + "\ M ns. haltP M ns \ {(\tp. tp = ([Bk], <(code M, ns)>))} H {(\tp. \k. tp = (Bk \ k, <0::nat>))}" and nh_case: - "\ M ns. \ haltP M ns \ {(\tp. tp = ([Bk], ))} H {(\tp. \k. tp = (Bk \ k, <1::nat>))}" + "\ M ns. \ haltP M ns \ {(\tp. tp = ([Bk], <(code M, ns)>))} H {(\tp. \k. tp = (Bk \ k, <1::nat>))}" begin (* invariants for H *) -abbreviation - "pre_H_inv M n \ \tp. tp = ([Bk], <[code M, n]>)" +abbreviation (input) + "pre_H_inv M ns \ \tp. tp = ([Bk], <(code M, ns::nat list)>)" -abbreviation +abbreviation (input) "post_H_halt_inv \ \tp. \k. tp = (Bk \ k, <1::nat>)" -abbreviation +abbreviation (input) "post_H_unhalt_inv \ \tp. \k. tp = (Bk \ k, <0::nat>)" lemma H_halt_inv: - assumes "\ haltP M [n]" - shows "{pre_H_inv M n} H {post_H_halt_inv}" + assumes "\ haltP M ns" + shows "{pre_H_inv M ns} H {post_H_halt_inv}" using assms nh_case by auto lemma H_unhalt_inv: - assumes "haltP M [n]" - shows "{pre_H_inv M n} H {post_H_unhalt_inv}" + assumes "haltP M ns" + shows "{pre_H_inv M ns} H {post_H_unhalt_inv}" using assms h_case by auto (* TM that produces the contradiction and its code *) @@ -1083,8 +1058,8 @@ shows "False" proof - (* invariants *) - def P1 \ "\tp. tp = ([]::cell list, <[code_tcontra]>)" - def P2 \ "\tp. tp = ([Bk], <[code_tcontra, code_tcontra]>)" + def P1 \ "\tp. tp = ([]::cell list, )" + def P2 \ "\tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)" def P3 \ "\tp. \k. tp = (Bk \ k, <1::nat>)" (* @@ -1106,8 +1081,9 @@ next case B_halt (* of H *) show "{P2} H {P3}" - unfolding P2_def P3_def - by (rule H_halt_inv[OF assms]) + unfolding P2_def P3_def + using H_halt_inv[OF assms] + by (simp add: tape_of_nat_pair tape_of_nl_abv) qed (simp) (* {P3} dither {P3} *) @@ -1123,10 +1099,10 @@ unfolding P1_def P3_def unfolding haltP_def unfolding Hoare_halt_def - apply(auto) + apply(auto) apply(drule_tac x = n in spec) - apply(case_tac "steps0 (Suc 0, [], <[code tcontra]>) tcontra n") - apply(auto) + apply(case_tac "steps0 (Suc 0, [], ) tcontra n") + apply(auto simp add: tape_of_nl_abv) done qed @@ -1136,8 +1112,8 @@ shows "False" proof - (* invariants *) - def P1 \ "\tp. tp = ([]::cell list, <[code_tcontra]>)" - def P2 \ "\tp. tp = ([Bk], <[code_tcontra, code_tcontra]>)" + def P1 \ "\tp. tp = ([]::cell list, )" + def P2 \ "\tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)" def Q3 \ "\tp. \k. tp = (Bk \ k, <0::nat>)" (* @@ -1159,8 +1135,8 @@ next case B_halt (* of H *) then show "{P2} H {Q3}" - unfolding P2_def Q3_def - by (rule H_unhalt_inv[OF assms]) + unfolding P2_def Q3_def using H_unhalt_inv[OF assms] + by(simp add: tape_of_nat_pair tape_of_nl_abv) qed (simp) (* {P3} dither loops *) @@ -1176,7 +1152,7 @@ unfolding P1_def unfolding haltP_def unfolding Hoare_halt_def Hoare_unhalt_def - by auto + by (auto simp add: tape_of_nl_abv) qed text {*