author | kleing |
Mon, 20 Aug 2007 00:22:18 +0200 | |
changeset 24332 | e3a2b75b1cf9 |
parent 24281 | 7d0334b69711 |
child 24530 | 1bac25879117 |
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 |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
9 |
Boolean_Algebra |
21256 | 10 |
Char_ord |
11 |
Coinductive_List |
|
12 |
Commutative_Ring |
|
15131 | 13 |
Continuity |
23854 | 14 |
Efficient_Nat |
22519 | 15 |
Eval |
24281 | 16 |
Eval_Witness |
23854 | 17 |
Executable_Rat |
22981
cf071f3fc4ae
A verified theory for rational numbers representation and simple calculations;
chaieb
parents:
22799
diff
changeset
|
18 |
Executable_Real |
23854 | 19 |
Executable_Set |
21256 | 20 |
FuncSet |
21 |
GCD |
|
22 |
Infinite_Set |
|
23854 | 23 |
ML_String |
15131 | 24 |
Multiset |
25 |
NatPair |
|
26 |
Nat_Infinity |
|
27 |
Nested_Environment |
|
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
28 |
Numeral_Type |
15470 | 29 |
OptionalSugar |
21256 | 30 |
Parity |
15131 | 31 |
Permutation |
22799
ed7d53db2170
moved code generation pretty integers and characters to separate theories
haftmann
parents:
22519
diff
changeset
|
32 |
Pretty_Char_chr |
ed7d53db2170
moved code generation pretty integers and characters to separate theories
haftmann
parents:
22519
diff
changeset
|
33 |
Pretty_Int |
15131 | 34 |
Primes |
35 |
Quotient |
|
21635
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents:
21256
diff
changeset
|
36 |
Ramsey |
21256 | 37 |
State_Monad |
22359
94a794672c8b
Added formalization of size-change principle (experimental).
krauss
parents:
21635
diff
changeset
|
38 |
Size_Change_Termination |
15131 | 39 |
While_Combinator |
40 |
Word |
|
41 |
Zorn |
|
42 |
begin |
|
10253 | 43 |
end |
44 |
(*>*) |