# HG changeset patch # User haftmann # Date 1199863929 -3600 # Node ID a6a21adf3b55dbd56f2c5eade101e8cd64550a92 # Parent d49bf150c9252f2299314f5026f86c0d9a34f966 tuned diff -r d49bf150c925 -r a6a21adf3b55 doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy Wed Jan 09 08:04:40 2008 +0100 +++ b/doc-src/IsarAdvanced/Codegen/Thy/Codegen.thy Wed Jan 09 08:32:09 2008 +0100 @@ -111,7 +111,7 @@ We define @{text find} and @{text update} functions: *} -fun +primrec find :: "('a\linorder, 'b) searchtree \ 'a \ 'b option" where "find (Leaf key val) it = (if it = key then Some val else None)" | "find (Branch t1 key t2) it = (if it \ key then find t1 it else find t2 it)" @@ -215,7 +215,7 @@ most cases code generation proceeds without further ado: *} -fun +primrec fac :: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = Suc n * fac n" @@ -262,6 +262,7 @@ pick_some :: "'a list \ 'a" where "pick_some = hd" (*>*) + export_code pick_some in SML file "examples/fail_const.ML" text {* \noindent will fail. *} @@ -282,18 +283,16 @@ The typical @{text HOL} tools are already set up in a way that function definitions introduced by @{text "\"}, - @{text "\"}, - @{text "\"}, @{text "\"}, + @{text "\"}, @{text "\"}, + @{text "\"}, @{text "\"}, @{text "\"} are implicitly propagated to this defining equation table. Specific theorems may be selected using an attribute: \emph{code func}. As example, a weight selector function: *} -consts - pick :: "(nat \ 'a) list \ nat \ 'a" - primrec + pick :: "(nat \ 'a) list \ nat \ 'a" where "pick (x#xs) n = (let (k, v) = x in if n < k then v else pick xs (n - k))" @@ -365,9 +364,8 @@ class null = type + fixes null :: 'a -fun - head :: "'a\null list \ 'a" -where +primrec + head :: "'a\null list \ 'a" where "head [] = null" | "head (x#xs) = x" @@ -586,7 +584,7 @@ by the code generator: *} -fun +primrec collect_duplicates :: "'a list \ 'a list \ 'a list \ 'a list" where "collect_duplicates xs ys [] = xs" | "collect_duplicates xs ys (z#zs) = (if z \ set xs @@ -636,15 +634,20 @@ hide (open) const "op =" setup {* Sign.parent_path *} (*>*) -instance * :: (ord, ord) ord - less_prod_def: - "p1 < p2 \ let (x1 \ 'a\ord, y1 \ 'b\ord) = p1; (x2, y2) = p2 in - x1 < x2 \ (x1 = x2 \ y1 < y2)" - less_eq_prod_def: - "p1 \ p2 \ let (x1 \ 'a\ord, y1 \ 'b\ord) = p1; (x2, y2) = p2 in - x1 < x2 \ (x1 = x2 \ y1 \ y2)" .. +instantiation * :: (ord, ord) ord +begin + +definition + [code func del]: "p1 < p2 \ (let (x1, y1) = p1; (x2, y2) = p2 in + x1 < x2 \ (x1 = x2 \ y1 < y2))" -lemmas [code func del] = less_prod_def less_eq_prod_def +definition + [code func del]: "p1 \ p2 \ (let (x1, y1) = p1; (x2, y2) = p2 in + x1 < x2 \ (x1 = x2 \ y1 \ y2))" + +instance .. + +end lemma ord_prod [code func(*<*), code func del(*>*)]: "(x1 \ 'a\ord, y1 \ 'b\ord) < (x2, y2) \ x1 < x2 \ (x1 = x2 \ y1 < y2)" @@ -809,7 +812,7 @@ SML code: *} -fun +primrec in_interval :: "nat \ nat \ nat \ bool" where "in_interval (k, l) n \ k \ n \ n \ l" (*<*) diff -r d49bf150c925 -r a6a21adf3b55 doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex --- a/doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex Wed Jan 09 08:04:40 2008 +0100 +++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex Wed Jan 09 08:32:09 2008 +0100 @@ -138,7 +138,7 @@ We define \isa{find} and \isa{update} functions:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{fun}\isamarkupfalse% +\isacommand{primrec}\isamarkupfalse% \isanewline \ \ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isasymColon}linorder{\isacharcomma}\ {\isacharprime}b{\isacharparenright}\ searchtree\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}find\ {\isacharparenleft}Leaf\ key\ val{\isacharparenright}\ it\ {\isacharequal}\ {\isacharparenleft}if\ it\ {\isacharequal}\ key\ then\ Some\ val\ else\ None{\isacharparenright}{\isachardoublequoteclose}\isanewline @@ -248,7 +248,7 @@ most cases code generation proceeds without further ado:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{fun}\isamarkupfalse% +\isacommand{primrec}\isamarkupfalse% \isanewline \ \ fac\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ \ \ {\isachardoublequoteopen}fac\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline @@ -310,6 +310,7 @@ \isadelimML % \endisadelimML +\isanewline \isacommand{export{\isacharunderscore}code}\isamarkupfalse% \ pick{\isacharunderscore}some\ \isakeyword{in}\ SML\ \isakeyword{file}\ {\isachardoublequoteopen}examples{\isacharslash}fail{\isacharunderscore}const{\isachardot}ML{\isachardoublequoteclose}% \begin{isamarkuptext}% @@ -335,20 +336,17 @@ The typical \isa{HOL} tools are already set up in a way that function definitions introduced by \isa{{\isasymDEFINITION}}, - \isa{{\isasymFUN}}, - \isa{{\isasymFUNCTION}}, \isa{{\isasymPRIMREC}}, + \isa{{\isasymPRIMREC}}, \isa{{\isasymFUN}}, + \isa{{\isasymFUNCTION}}, \isa{{\isasymCONSTDEFS}}, \isa{{\isasymRECDEF}} are implicitly propagated to this defining equation table. Specific theorems may be selected using an attribute: \emph{code func}. As example, a weight selector function:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{consts}\isamarkupfalse% -\isanewline -\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\isanewline -\isanewline \isacommand{primrec}\isamarkupfalse% \isanewline +\ \ pick\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}nat\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ list\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}pick\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ n\ {\isacharequal}\ {\isacharparenleft}let\ {\isacharparenleft}k{\isacharcomma}\ v{\isacharparenright}\ {\isacharequal}\ x\ in\isanewline \ \ \ \ if\ n\ {\isacharless}\ k\ then\ v\ else\ pick\ xs\ {\isacharparenleft}n\ {\isacharminus}\ k{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}% \begin{isamarkuptext}% @@ -454,10 +452,9 @@ \ null\ {\isacharequal}\ type\ {\isacharplus}\isanewline \ \ \isakeyword{fixes}\ null\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline \isanewline -\isacommand{fun}\isamarkupfalse% +\isacommand{primrec}\isamarkupfalse% \isanewline -\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\isanewline -\isakeyword{where}\isanewline +\ \ head\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a{\isasymColon}null\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}head\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ null{\isachardoublequoteclose}\isanewline \ \ {\isacharbar}\ {\isachardoublequoteopen}head\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequoteclose}% \begin{isamarkuptext}% @@ -740,7 +737,7 @@ by the code generator:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{fun}\isamarkupfalse% +\isacommand{primrec}\isamarkupfalse% \isanewline \ \ collect{\isacharunderscore}duplicates\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ \ \ {\isachardoublequoteopen}collect{\isacharunderscore}duplicates\ xs\ ys\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline @@ -810,14 +807,22 @@ \isadelimML % \endisadelimML -\isacommand{instance}\isamarkupfalse% +\isacommand{instantiation}\isamarkupfalse% \ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ord{\isacharcomma}\ ord{\isacharparenright}\ ord\isanewline -\ \ less{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymequiv}\ let\ {\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline -\ \ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\isanewline -\ \ \ \ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymequiv}\ let\ {\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}ord{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}ord{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline -\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% +\isakeyword{begin}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isacharless}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{definition}\isamarkupfalse% +\isanewline +\ \ {\isacharbrackleft}code\ func\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}p{\isadigit{1}}\ {\isasymle}\ p{\isadigit{2}}\ {\isasymlongleftrightarrow}\ {\isacharparenleft}let\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{1}}{\isacharsemicolon}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ p{\isadigit{2}}\ in\isanewline +\ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ {\isacharparenleft}x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\isacommand{instance}\isamarkupfalse% +% \isadelimproof \ % \endisadelimproof @@ -833,8 +838,8 @@ \endisadelimproof \isanewline \isanewline -\isacommand{lemmas}\isamarkupfalse% -\ {\isacharbrackleft}code\ func\ del{\isacharbrackright}\ {\isacharequal}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\isanewline +\isacommand{end}\isamarkupfalse% +\isanewline \isanewline \isacommand{lemma}\isamarkupfalse% \ ord{\isacharunderscore}prod\ {\isacharbrackleft}code\ func{\isacharbrackright}{\isacharcolon}\isanewline @@ -1151,7 +1156,7 @@ SML code:% \end{isamarkuptext}% \isamarkuptrue% -\isacommand{fun}\isamarkupfalse% +\isacommand{primrec}\isamarkupfalse% \isanewline \ \ in{\isacharunderscore}interval\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}in{\isacharunderscore}interval\ {\isacharparenleft}k{\isacharcomma}\ l{\isacharparenright}\ n\ {\isasymlongleftrightarrow}\ k\ {\isasymle}\ n\ {\isasymand}\ n\ {\isasymle}\ l{\isachardoublequoteclose}% diff -r d49bf150c925 -r a6a21adf3b55 doc-src/IsarAdvanced/Codegen/codegen.tex --- a/doc-src/IsarAdvanced/Codegen/codegen.tex Wed Jan 09 08:04:40 2008 +0100 +++ b/doc-src/IsarAdvanced/Codegen/codegen.tex Wed Jan 09 08:32:09 2008 +0100 @@ -50,6 +50,7 @@ \newcommand{\isasymFUN}{\cmd{fun}} \newcommand{\isasymFUNCTION}{\cmd{function}} \newcommand{\isasymPRIMREC}{\cmd{primrec}} +\newcommand{\isasymCONSTDEFS}{\cmd{constdefs}} \newcommand{\isasymRECDEF}{\cmd{recdef}} \isakeeptag{tt}