# HG changeset patch # User Christian Urban # Date 1359164208 0 # Node ID 4c8325c64dab129a4c7e2b02acbc6d473880b67a # Parent 8dc659af1bd232a6f4c15529ce7871d956352c9f updated paper diff -r 8dc659af1bd2 -r 4c8325c64dab Paper/Paper.thy --- a/Paper/Paper.thy Fri Jan 25 22:12:01 2013 +0000 +++ b/Paper/Paper.thy Sat Jan 26 01:36:48 2013 +0000 @@ -531,16 +531,19 @@ We often need to restrict tapes to be in \emph{standard form}, which means the left list of the tape is either empty or only contains @{text "Bk"}s, and the right list contains some ``clusters'' of @{text "Oc"}s separted by single - blanks and can be followed by some blanks. To make this formal we define the - following function + blanks. To make this formal we define the following function \begin{center} - foo + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} + @{thm (lhs) tape_of_nat_list.simps(1)} & @{text "\"} & @{thm (rhs) tape_of_nat_list.simps(1)}\\ + @{thm (lhs) tape_of_nat_list.simps(2)} & @{text "\"} & @{thm (rhs) tape_of_nat_list.simps(2)}\\ + @{thm (lhs) tape_of_nat_list.simps(3)} & @{text "\"} & @{thm (rhs) tape_of_nat_list.simps(3)} + \end{tabular} \end{center} \noindent - A standard tape is then of the form @{text "(Bk\<^isup>k,\n\<^isub>1,...,n\<^isub>m\ @ Bk\<^isup>l)"} for some @{text k}, - @{text l} and @{text "n\<^isub>i"}. Note that the head in a standard tape ``points'' to the + A standard tape is then of the form @{text "(Bk\<^isup>l,\n\<^isub>1,...,n\<^isub>m\)"} for some @{text l} + and @{text "n\<^isub>i"}. Note that the head in a standard tape ``points'' to the leftmost @{term "Oc"} on the tape. diff -r 8dc659af1bd2 -r 4c8325c64dab paper.pdf Binary file paper.pdf has changed diff -r 8dc659af1bd2 -r 4c8325c64dab thys/turing_basic.thy --- a/thys/turing_basic.thy Fri Jan 25 22:12:01 2013 +0000 +++ b/thys/turing_basic.thy Sat Jan 26 01:36:48 2013 +0000 @@ -194,13 +194,15 @@ "tape_of_nat_list [n] = Oc\(Suc n)" | "tape_of_nat_list (n#ns) = Oc\(Suc n) @ Bk # (tape_of_nat_list ns)" -defs (overloaded) - tape_of_nl_abv: " \ tape_of_nat_list am" - tape_of_nat_abv : "<(n::nat)> \ Oc\(Suc n)" +fun tape_of_nat_pair :: "nat \ nat \ cell list" + where + "tape_of_nat_pair (n, m) = Oc\(Suc n) @ [Bk] @ Oc\(Suc m)" + -definition tinres :: "cell list \ cell list \ bool" - where - "tinres xs ys = (\n. xs = ys @ Bk \ n \ ys = xs @ Bk \ n)" +defs (overloaded) + tape_of_nl_abv: " \ tape_of_nat_list ns" + tape_of_nat_abv: "<(n::nat)> \ Oc\(Suc n)" + tape_of_nat_pair: "

\ tape_of_nat_pair p" fun shift :: "instr list \ nat \ instr list" diff -r 8dc659af1bd2 -r 4c8325c64dab thys/uncomputable.thy --- a/thys/uncomputable.thy Fri Jan 25 22:12:01 2013 +0000 +++ b/thys/uncomputable.thy Sat Jan 26 01:36:48 2013 +0000 @@ -204,16 +204,16 @@ 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 = Oc\i \ r = 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 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_loop5_loop x (l, r) = - (\ i j k t. i + j = Suc x \ i > 0 \ j > 0 \ k + t = j \ t > 0 \ l = Oc\k@Bk\j@Oc\i \ r = Oc\t)" + (\ 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) = - (\ i j. i + j = Suc x \ i > 0 \ j > 0 \ l = Bk\(j - 1)@Oc\i \ r = Bk # Oc\j)" + (\ i j. i + j = Suc x \ i > 0 \ j > 0 \ (l, r) = (Bk\(j - 1)@Oc\i, Bk # Oc\j))" | "inv_loop6_loop x (l, r) = - (\ i j k t. i + j = Suc x \ i > 0 \ k + t + 1 = j \ l = Bk\k @ Oc\i \ r = Bk\(Suc t) @ Oc\j)" + (\ i j k t. i + j = Suc x \ i > 0 \ k + t + 1 = j \ (l, r) = (Bk\k @ Oc\i, Bk\(Suc t) @ Oc\j))" | "inv_loop6_exit x (l, r) = - (\ i j. i + j = x \ j > 0 \ l = Oc\i \ r = Oc # Bk\j @ Oc\j)" + (\ i j. i + j = x \ j > 0 \ (l, r) = (Oc\i, Oc#Bk\j @ Oc\j))" fun inv_loop0 :: "nat \ tape \ bool" and @@ -224,13 +224,13 @@ inv_loop5 :: "nat \ tape \ bool" and inv_loop6 :: "nat \ tape \ bool" where - "inv_loop0 x (l, r) = (l = [Bk] \ r = Oc # Bk\x @ Oc\x \ x > 0 )" + "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 = Oc\i \ r = any#Bk\j@Oc\j)" +| "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 = Bk\k@Oc\i \ r = Bk\t@Oc\j)" + (\ 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 = Oc\k @ Bk\(Suc j)@Oc\i \ r = Oc\t)" + (\ 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))" @@ -1026,7 +1026,7 @@ definition haltP :: "tprog \ nat list \ bool" where - "haltP p lm \ \n a b c. steps (1, [], ) p n = (0, Bk \ a, Oc \ b @ Bk \ c)" + "haltP p lm \ \n a m. steps (1, [], ) p n = (0, Bk \ a, )" abbreviation "haltP0 p lm \ haltP (p, 0) lm" @@ -1057,9 +1057,9 @@ The following two assumptions specifies that @{text "H"} does solve the Halting problem. *} and h_case: - "\ M lm. (haltP0 M lm) \ \ na nb. (steps0 (1, [Bk], ) H na = (0, Bk \ nb, [Oc]))" + "\ M lm. (haltP0 M lm) \ \stp n. (steps0 (1, [Bk], ) H stp = (0, Bk \ n, [Oc]))" and nh_case: - "\ M lm. (\ haltP0 M lm) \ \ na nb. (steps0 (1, [Bk], ) H na = (0, Bk \ nb, [Oc, Oc]))" + "\ M lm. (\ haltP0 M lm) \ \stp n. (steps0 (1, [Bk], ) H stp = (0, Bk \ n, [Oc, Oc]))" begin (* invariants for H *) @@ -1079,7 +1079,7 @@ proof - obtain stp i where "steps0 (1, [Bk], <[code M, n]>) H stp = (0, Bk \ i, [Oc, Oc])" - using nh_case[of "M" "[n]", OF assms] by auto + using nh_case assms by blast then show "{pre_H_inv M n} H {post_H_halt_inv}" unfolding Hoare_halt_def apply(auto) @@ -1094,7 +1094,7 @@ proof - obtain stp i where "steps0 (1, [Bk], <[code M, n]>) H stp = (0, Bk \ i, [Oc])" - using h_case[of "M" "[n]", OF assms] by auto + using h_case assms by blast then show "{pre_H_inv M n} H {post_H_unhalt_inv}" unfolding Hoare_halt_def apply(auto) @@ -1156,11 +1156,11 @@ unfolding haltP_def unfolding Hoare_halt_def apply(auto) - apply(erule_tac x = n in allE) + apply(drule_tac x = n in spec) apply(case_tac "steps0 (Suc 0, [], <[code tcontra]>) tcontra n") - apply(simp, auto) - apply(erule_tac x = 2 in allE, erule_tac x = 0 in allE) - apply(simp add: numeral) + apply(auto) + apply(drule_tac x = 1 in spec) + apply(simp add: numeral tape_of_nat_abv) done qed