21917
|
1 |
(* ID: $Id$
|
|
2 |
Author: Florian Haftmann, TU Muenchen
|
|
3 |
*)
|
|
4 |
|
|
5 |
header {* A huge set of executable constants *}
|
|
6 |
|
|
7 |
theory ExecutableContent
|
|
8 |
imports
|
|
9 |
Main
|
22521
|
10 |
Eval
|
23690
|
11 |
"~~/src/HOL/ex/Records"
|
21917
|
12 |
AssocList
|
|
13 |
Binomial
|
|
14 |
Commutative_Ring
|
23690
|
15 |
"~~/src/HOL/ex/Commutative_Ring_Complete"
|
24197
|
16 |
Executable_Rat
|
23016
|
17 |
Executable_Real
|
21917
|
18 |
GCD
|
|
19 |
List_Prefix
|
|
20 |
Nat_Infinity
|
|
21 |
NatPair
|
22195
|
22 |
Nested_Environment
|
21917
|
23 |
Permutation
|
|
24 |
Primes
|
|
25 |
Product_ord
|
|
26 |
SetsAndFunctions
|
|
27 |
State_Monad
|
|
28 |
While_Combinator
|
|
29 |
Word
|
|
30 |
begin
|
|
31 |
|
|
32 |
definition
|
|
33 |
n :: nat where
|
|
34 |
"n = 42"
|
|
35 |
|
|
36 |
definition
|
|
37 |
k :: "int" where
|
|
38 |
"k = -42"
|
|
39 |
|
|
40 |
datatype mut1 = Tip | Top mut2
|
|
41 |
and mut2 = Tip | Top mut1
|
|
42 |
|
|
43 |
consts
|
|
44 |
mut1 :: "mut1 \<Rightarrow> mut1"
|
|
45 |
mut2 :: "mut2 \<Rightarrow> mut2"
|
|
46 |
|
|
47 |
primrec
|
|
48 |
"mut1 mut1.Tip = mut1.Tip"
|
|
49 |
"mut1 (mut1.Top x) = mut1.Top (mut2 x)"
|
|
50 |
"mut2 mut2.Tip = mut2.Tip"
|
|
51 |
"mut2 (mut2.Top x) = mut2.Top (mut1 x)"
|
|
52 |
|
|
53 |
definition
|
|
54 |
"mystring = ''my home is my castle''"
|
|
55 |
|
|
56 |
text {* nested lets and such *}
|
|
57 |
|
|
58 |
definition
|
|
59 |
"abs_let x = (let (y, z) = x in (\<lambda>u. case u of () \<Rightarrow> (y + y)))"
|
|
60 |
|
|
61 |
definition
|
|
62 |
"nested_let x = (let (y, z) = x in let w = y z in w * w)"
|
|
63 |
|
|
64 |
definition
|
|
65 |
"case_let x = (let (y, z) = x in case y of () => z)"
|
|
66 |
|
|
67 |
definition
|
|
68 |
"base_case f = f list_case"
|
|
69 |
|
|
70 |
definition
|
|
71 |
"apply_tower = (\<lambda>x. x (\<lambda>x. x (\<lambda>x. x)))"
|
|
72 |
|
|
73 |
definition
|
|
74 |
"keywords fun datatype x instance funa classa =
|
|
75 |
Suc fun + datatype * x mod instance - funa - classa"
|
|
76 |
|
|
77 |
hide (open) const keywords
|
|
78 |
|
|
79 |
definition
|
|
80 |
"shadow keywords = keywords @ [ExecutableContent.keywords 0 0 0 0 0 0]"
|
|
81 |
|
24197
|
82 |
definition
|
|
83 |
foo :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> rat" where
|
|
84 |
"foo r s t = (t + s) / t"
|
|
85 |
|
|
86 |
definition
|
|
87 |
bar :: "rat \<Rightarrow> rat \<Rightarrow> rat \<Rightarrow> bool" where
|
|
88 |
"bar r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"
|
|
89 |
|
|
90 |
definition
|
|
91 |
"R1 = Fract 3 7"
|
|
92 |
|
|
93 |
definition
|
|
94 |
"R2 = Fract (-7) 5"
|
|
95 |
|
|
96 |
definition
|
|
97 |
"R3 = Fract 11 (-9)"
|
|
98 |
|
|
99 |
definition
|
|
100 |
"foobar = (foo R1 1 R3, bar R2 0 R3, foo R1 R3 R2)"
|
|
101 |
|
|
102 |
definition
|
|
103 |
foo' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
|
|
104 |
"foo' r s t = (t + s) / t"
|
|
105 |
|
|
106 |
definition
|
|
107 |
bar' :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
|
|
108 |
"bar' r s t \<longleftrightarrow> (r - s) \<le> t \<or> (s - t) \<le> r"
|
|
109 |
|
|
110 |
definition
|
|
111 |
"R1' = real_of_rat (Fract 3 7)"
|
|
112 |
|
|
113 |
definition
|
|
114 |
"R2' = real_of_rat (Fract (-7) 5)"
|
|
115 |
|
|
116 |
definition
|
|
117 |
"R3' = real_of_rat (Fract 11 (-9))"
|
|
118 |
|
|
119 |
definition
|
|
120 |
"foobar' = (foo' R1' 1 R3', bar' R2' 0 R3', foo' R1' R3' R2')"
|
|
121 |
|
21917
|
122 |
end
|