| author | wenzelm |
| Fri, 17 Aug 2007 00:03:50 +0200 | |
| changeset 24300 | e170cee91c66 |
| parent 24281 | 7d0334b69711 |
| child 24332 | e3a2b75b1cf9 |
| permissions | -rw-r--r-- |
| 20809 | 1 |
(* $Id$ *) |
| 10253 | 2 |
(*<*) |
| 15131 | 3 |
theory Library |
| 15140 | 4 |
imports |
| 24197 | 5 |
Abstract_Rat |
| 21256 | 6 |
AssocList |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
16109
diff
changeset
|
7 |
BigO |
| 21256 | 8 |
Binomial |
9 |
Char_ord |
|
10 |
Coinductive_List |
|
11 |
Commutative_Ring |
|
| 15131 | 12 |
Continuity |
| 23854 | 13 |
Efficient_Nat |
| 22519 | 14 |
Eval |
| 24281 | 15 |
Eval_Witness |
| 23854 | 16 |
Executable_Rat |
|
22981
cf071f3fc4ae
A verified theory for rational numbers representation and simple calculations;
chaieb
parents:
22799
diff
changeset
|
17 |
Executable_Real |
| 23854 | 18 |
Executable_Set |
| 21256 | 19 |
FuncSet |
20 |
GCD |
|
21 |
Infinite_Set |
|
| 23854 | 22 |
ML_String |
| 15131 | 23 |
Multiset |
24 |
NatPair |
|
25 |
Nat_Infinity |
|
26 |
Nested_Environment |
|
| 15470 | 27 |
OptionalSugar |
| 21256 | 28 |
Parity |
| 15131 | 29 |
Permutation |
|
22799
ed7d53db2170
moved code generation pretty integers and characters to separate theories
haftmann
parents:
22519
diff
changeset
|
30 |
Pretty_Char_chr |
|
ed7d53db2170
moved code generation pretty integers and characters to separate theories
haftmann
parents:
22519
diff
changeset
|
31 |
Pretty_Int |
| 15131 | 32 |
Primes |
33 |
Quotient |
|
|
21635
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents:
21256
diff
changeset
|
34 |
Ramsey |
| 21256 | 35 |
State_Monad |
|
22359
94a794672c8b
Added formalization of size-change principle (experimental).
krauss
parents:
21635
diff
changeset
|
36 |
Size_Change_Termination |
| 15131 | 37 |
While_Combinator |
38 |
Word |
|
39 |
Zorn |
|
40 |
begin |
|
| 10253 | 41 |
end |
42 |
(*>*) |