| author | chaieb |
| Wed, 27 Feb 2008 14:39:58 +0100 | |
| changeset 26161 | 34cb0b457dcc |
| parent 26030 | 4ae4ea600e8f |
| child 26447 | fef9dde61a46 |
| permissions | -rw-r--r-- |
| 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 |
|
| 25536 | 10 |
Eval |
| 25963 | 11 |
Code_Index |
| 26030 | 12 |
"~~/src/HOL/ex/Records" |
| 21917 | 13 |
AssocList |
14 |
Binomial |
|
15 |
Commutative_Ring |
|
| 23690 | 16 |
"~~/src/HOL/ex/Commutative_Ring_Complete" |
|
24530
1bac25879117
Integrated Executable_Rat and Executable_Real theories into
berghofe
parents:
24433
diff
changeset
|
17 |
"~~/src/HOL/Real/RealDef" |
| 21917 | 18 |
GCD |
19 |
List_Prefix |
|
20 |
Nat_Infinity |
|
21 |
NatPair |
|
|
24433
4a405457e9d6
added explicit equation for equality of nested environments
haftmann
parents:
24423
diff
changeset
|
22 |
Nested_Environment |
| 21917 | 23 |
Permutation |
24 |
Primes |
|
25 |
Product_ord |
|
26 |
SetsAndFunctions |
|
27 |
State_Monad |
|
28 |
While_Combinator |
|
29 |
Word |
|
30 |
begin |
|
31 |
||
| 26022 | 32 |
lemma [code func, code func del]: "(Eval.term_of \<Colon> index \<Rightarrow> term) = Eval.term_of" .. |
| 25963 | 33 |
declare fast_bv_to_nat_helper.simps [code func del] |
| 24197 | 34 |
|
| 21917 | 35 |
end |