src/HOL/ex/ExecutableContent.thy
changeset 25963 07e08dad8a77
parent 25572 0c9052719f20
child 26022 b30a342a6e29
--- a/src/HOL/ex/ExecutableContent.thy	Fri Jan 25 14:53:55 2008 +0100
+++ b/src/HOL/ex/ExecutableContent.thy	Fri Jan 25 14:53:56 2008 +0100
@@ -8,6 +8,7 @@
 imports
   Main
   Eval
+  Code_Index
   "~~/src/HOL/ex/Records"
   AssocList
   Binomial
@@ -28,93 +29,7 @@
   Word
 begin
 
-definition
-  n :: nat where
-  "n = 42"
-
-definition
-  k :: "int" where
-  "k = -42"
-
-datatype mut1 = Tip | Top mut2
-  and mut2 = Tip | Top mut1
-
-primrec
-  mut1 :: "mut1 \<Rightarrow> mut1"
-  and mut2 :: "mut2 \<Rightarrow> mut2"
-where
-  "mut1 mut1.Tip = mut1.Tip"
-  | "mut1 (mut1.Top x) = mut1.Top (mut2 x)"
-  | "mut2 mut2.Tip = mut2.Tip"
-  | "mut2 (mut2.Top x) = mut2.Top (mut1 x)"
-
-definition
-  "mystring = ''my home is my castle''"
-
-text {* nested lets and such *}
-
-definition
-  "abs_let x = (let (y, z) = x in (\<lambda>u. case u of () \<Rightarrow> (y + y)))"
-
-definition
-  "nested_let x = (let (y, z) = x in let w = y z in w * w)"
-
-definition
-  "case_let x = (let (y, z) = x in case y of () => z)"
-
-definition
-  "base_case f = f list_case"
-
-definition
-  "apply_tower = (\<lambda>x. x (\<lambda>x. x (\<lambda>x. x)))"
-
-definition
-  "keywords fun datatype x instance funa classa =
-    Suc fun + datatype * x mod instance - funa - classa"
-
-hide (open) const keywords
-
-definition
-  "shadow keywords = keywords @ [ExecutableContent.keywords 0 0 0 0 0 0]"
-
-definition
-  foo :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> rat" where
-  "foo r s t = (t + s) / t"
-
-definition
-  bar :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> bool" where
-  "bar r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"
-
-definition
-  "R1 = Fract 3 7"
-
-definition
-  "R2 = Fract (-7) 5"
-
-definition
-  "R3 = Fract 11 (-9)"
-
-definition
-  "foobar = (foo R1 1 R3, bar R2 0 R3, foo R1 R3 R2)"
-
-definition
-  foo' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
-  "foo' r s t = (t + s) / t"
-
-definition
-  bar' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
-  "bar' r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"
-
-definition
-  "R1' = real_of_rat (Fract 3 7)"
-
-definition
-  "R2' = real_of_rat (Fract (-7) 5)"
-
-definition
-  "R3' = real_of_rat (Fract 11 (-9))"
-
-definition
-  "foobar' = (foo' R1' 1 R3', bar' R2' 0 R3', foo' R1' R3' R2')"
+declare term_of_index.simps [code func del]
+declare fast_bv_to_nat_helper.simps [code func del]
 
 end