| author | nipkow |
| Tue, 18 Sep 2007 05:42:46 +0200 | |
| changeset 24615 | 17dbd993293d |
| parent 24530 | 1bac25879117 |
| child 24626 | 85eceef2edc7 |
| 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_Set |
| 21256 | 18 |
FuncSet |
19 |
GCD |
|
20 |
Infinite_Set |
|
| 23854 | 21 |
ML_String |
| 15131 | 22 |
Multiset |
23 |
NatPair |
|
24 |
Nat_Infinity |
|
25 |
Nested_Environment |
|
|
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
26 |
Numeral_Type |
| 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 |
| 24615 | 33 |
Quicksort |
| 15131 | 34 |
Quotient |
|
21635
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents:
21256
diff
changeset
|
35 |
Ramsey |
| 21256 | 36 |
State_Monad |
|
22359
94a794672c8b
Added formalization of size-change principle (experimental).
krauss
parents:
21635
diff
changeset
|
37 |
Size_Change_Termination |
| 15131 | 38 |
While_Combinator |
39 |
Word |
|
40 |
Zorn |
|
41 |
begin |
|
| 10253 | 42 |
end |
43 |
(*>*) |