# HG changeset patch # User Christian Urban # Date 1359898274 0 # Node ID dfc629cd11de8303226935e92ddacdd636f1fc30 # Parent 480aae81b489325578f08bd911a307fb7779e416 made uncomputable compatible with abacus diff -r 480aae81b489 -r dfc629cd11de Paper/Paper.thy --- a/Paper/Paper.thy Sun Feb 03 12:24:28 2013 +0000 +++ b/Paper/Paper.thy Sun Feb 03 13:31:14 2013 +0000 @@ -1,6 +1,6 @@ (*<*) theory Paper -imports "../thys/uncomputable" +imports "../thys/abacus" begin (* @@ -1119,9 +1119,6 @@ to an contradiction, which means we have to abondon our assumption that there exists a Turing machine @{term H} which can decide whether Turing machines terminate. - - - *} @@ -1135,22 +1132,22 @@ over an unlimited number of registers $R_0$, $R_1$, \ldots each being able to hold an arbitrary large natural number. We use natural numbers to refer to registers, but also - to refer to \emph{opcodes} of abacus - machines. Obcodes are given by the datatype + to refer to \emph{statements} of abacus programs. Statements + are given by the datatype \begin{center} - \begin{tabular}{rcll} + \begin{tabular}{rcl@ {\hspace{10mm}}l} @{text "o"} & $::=$ & @{term "Inc R\"} & increment register $R$ by one\\ & $\mid$ & @{term "Dec R\ o\"} & if content of $R$ is non-zero,\\ & & & then decrement it by one\\ - & & & otherwise jump to opcode $o$\\ - & $\mid$ & @{term "Goto o\"} & jump to opcode $o$ + & & & otherwise jump to statement $o$\\ + & $\mid$ & @{term "Goto o\"} & jump to statement $o$ \end{tabular} \end{center} \noindent A \emph{program} of an abacus machine is a list of such - obcodes. For example the program clearing the register + statements. For example the program clearing the register $R$ (setting it to 0) can be defined as follows: \begin{center} diff -r 480aae81b489 -r dfc629cd11de Paper/ROOT.ML --- a/Paper/ROOT.ML Sun Feb 03 12:24:28 2013 +0000 +++ b/Paper/ROOT.ML Sun Feb 03 13:31:14 2013 +0000 @@ -1,8 +1,8 @@ no_document use_thys ["../thys/turing_basic", "../thys/turing_hoare", - "../thys/uncomputable"(*, - "../thys/abacus", + "../thys/uncomputable", + "../thys/abacus"(*, "../thys/rec_def", "../thys/recursive"*)]; diff -r 480aae81b489 -r dfc629cd11de paper.pdf Binary file paper.pdf has changed diff -r 480aae81b489 -r dfc629cd11de thys/abacus.thy --- a/thys/abacus.thy Sun Feb 03 12:24:28 2013 +0000 +++ b/thys/abacus.thy Sun Feb 03 13:31:14 2013 +0000 @@ -3,9 +3,19 @@ *} theory abacus -imports Main turing_basic +imports uncomputable begin +(* +declare tm_comp.simps [simp add] +declare adjust.simps[simp add] +declare shift.simps[simp add] +declare tm_wf.simps[simp add] +declare step.simps[simp add] +declare steps.simps[simp add] +*) +declare replicate_Suc[simp add] + text {* {\em Abacus} instructions: *} @@ -1300,8 +1310,7 @@ inv_locate_b.simps[simp del] lemma [intro]: "\rn. [Bk] = Bk \ rn" -apply(rule_tac x = "Suc 0" in exI, simp) -done +by (metis replicate_0 replicate_Suc) lemma [intro]: "at_begin_norm (as, am) (q, aaa, []) ires \ at_begin_norm (as, am) (q, aaa, [Bk]) ires" @@ -1877,7 +1886,6 @@ \ inv_after_write (as, lm) (y, l, Oc # r) ires" apply(auto simp: inv_on_right_moving.simps inv_after_write.simps ) apply(case_tac mr, auto simp: split: if_splits) -apply(case_tac [!] mr, simp_all) done lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires\