# HG changeset patch # User haftmann # Date 1233686886 -3600 # Node ID a342da8ddf3986ca639fdef9a389d1fd3f1e98af # Parent c78806b621e12e72ca9a2d92578f449a2bcd1519 adjusted theory name diff -r c78806b621e1 -r a342da8ddf39 doc-src/IsarAdvanced/Codegen/Thy/Further.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Further.thy Tue Feb 03 19:37:30 2009 +0100 +++ b/doc-src/IsarAdvanced/Codegen/Thy/Further.thy Tue Feb 03 19:48:06 2009 +0100 @@ -96,7 +96,7 @@ allows to use pattern matching on constructors stemming from compiled @{text datatypes}. - For a less simplistic example, theory @{theory ReflectedFerrack} is + For a less simplistic example, theory @{theory Ferrack} is a good reference. *} diff -r c78806b621e1 -r a342da8ddf39 doc-src/IsarAdvanced/Codegen/Thy/Program.thy --- a/doc-src/IsarAdvanced/Codegen/Thy/Program.thy Tue Feb 03 19:37:30 2009 +0100 +++ b/doc-src/IsarAdvanced/Codegen/Thy/Program.thy Tue Feb 03 19:48:06 2009 +0100 @@ -317,12 +317,6 @@ \end{itemize} *} -code_datatype %invisible "0::nat" Suc -declare %invisible plus_Dig [code del] -declare %invisible One_nat_def [code inline] -declare %invisible add_Suc_shift [code] -lemma %invisible [code]: "0 + n = (n \ nat)" by simp - subsection {* Equality and wellsortedness *} @@ -459,14 +453,14 @@ in the following example, again for amortised queues: *} -fun %quote strict_dequeue :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue (Queue xs (y # ys)) = (y, Queue xs ys)" - | "strict_dequeue (Queue xs []) = - (case rev xs of y # ys \ (y, Queue [] ys))" +fun %quote strict_dequeue :: "'a Introduction.queue \ 'a \ 'a Introduction.queue" where + "strict_dequeue (Introduction.Queue xs (y # ys)) = (y, Introduction.Queue xs ys)" + | "strict_dequeue (Introduction.Queue xs []) = + (case rev xs of y # ys \ (y, Introduction.Queue [] ys))" text {* \noindent In the corresponding code, there is no equation - for the pattern @{term "Queue [] []"}: + for the pattern @{term "Introduction.Queue [] []"}: *} text %quote {*@{code_stmts strict_dequeue (consts) strict_dequeue (Haskell)}*} @@ -478,15 +472,15 @@ axiomatization %quote empty_queue :: 'a -function %quote strict_dequeue' :: "'a queue \ 'a \ 'a queue" where - "strict_dequeue' (Queue xs []) = (if xs = [] then empty_queue - else strict_dequeue' (Queue [] (rev xs)))" - | "strict_dequeue' (Queue xs (y # ys)) = - (y, Queue xs ys)" +function %quote strict_dequeue' :: "'a Introduction.queue \ 'a \ 'a Introduction.queue" where + "strict_dequeue' (Introduction.Queue xs []) = (if xs = [] then empty_queue + else strict_dequeue' (Introduction.Queue [] (rev xs)))" + | "strict_dequeue' (Introduction.Queue xs (y # ys)) = + (y, Introduction.Queue xs ys)" by pat_completeness auto termination %quote strict_dequeue' -by (relation "measure (\q::'a queue. case q of Queue xs ys \ length xs)") simp_all +by (relation "measure (\q::'a Introduction.queue. case q of Introduction.Queue xs ys \ length xs)") simp_all text {* \noindent For technical reasons the definition of