| author | chaieb |
| Wed, 16 May 2007 09:45:22 +0200 | |
| changeset 22982 | bff3fcdeecd3 |
| parent 22981 | cf071f3fc4ae |
| child 23016 | fd7cd1edc18d |
| permissions | -rw-r--r-- |
| 22553 | 1 |
|
| 21917 | 2 |
(* ID: $Id$ |
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
*) |
|
5 |
||
6 |
header {* A huge set of executable constants *}
|
|
7 |
||
8 |
theory ExecutableContent |
|
9 |
imports |
|
10 |
Main |
|
| 22521 | 11 |
Eval |
| 21917 | 12 |
"~~/src/HOL/ex/Records" |
13 |
AssocList |
|
14 |
Binomial |
|
15 |
Commutative_Ring |
|
| 22748 | 16 |
"~~/src/HOL/ex/Commutative_Ring_Complete" |
|
22981
cf071f3fc4ae
A verified theory for rational numbers representation and simple calculations;
chaieb
parents:
22748
diff
changeset
|
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 |
||
82 |
end |