| author | nipkow |
| Sun, 15 Feb 2009 11:26:38 +0100 | |
| changeset 29920 | b95f5b8b93dd |
| parent 29879 | 4425849f5db7 |
| child 29985 | 57975b45ab70 |
| child 30240 | 5b25fee0362c |
| permissions | -rw-r--r-- |
| 10253 | 1 |
(*<*) |
| 15131 | 2 |
theory Library |
| 15140 | 3 |
imports |
|
27652
818666de6c24
refined code generator setup for rational numbers; more simplification rules for rational numbers
haftmann
parents:
27475
diff
changeset
|
4 |
Abstract_Rat |
| 21256 | 5 |
AssocList |
|
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
16109
diff
changeset
|
6 |
BigO |
| 21256 | 7 |
Binomial |
|
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
8 |
Boolean_Algebra |
| 21256 | 9 |
Char_ord |
| 26170 | 10 |
Code_Char_chr |
|
24994
c385c4eabb3b
consolidated naming conventions for code generator theories
haftmann
parents:
24626
diff
changeset
|
11 |
Code_Index |
| 26170 | 12 |
Code_Integer |
| 21256 | 13 |
Coinductive_List |
14 |
Commutative_Ring |
|
| 15131 | 15 |
Continuity |
| 29026 | 16 |
ContNotDenum |
| 26170 | 17 |
Countable |
| 29847 | 18 |
Determinants |
| 23854 | 19 |
Efficient_Nat |
| 26348 | 20 |
Enum |
| 24281 | 21 |
Eval_Witness |
| 23854 | 22 |
Executable_Set |
|
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28668
diff
changeset
|
23 |
Float |
| 29688 | 24 |
Formal_Power_Series |
| 21256 | 25 |
FuncSet |
| 29879 | 26 |
Fundamental_Theorem_Algebra |
|
27475
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
huffman
parents:
27368
diff
changeset
|
27 |
Infinite_Set |
| 26173 | 28 |
ListVector |
| 29708 | 29 |
Mapping |
| 15131 | 30 |
Multiset |
31 |
Nat_Infinity |
|
32 |
Nested_Environment |
|
|
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
33 |
Numeral_Type |
| 15470 | 34 |
OptionalSugar |
| 26232 | 35 |
Option_ord |
| 15131 | 36 |
Permutation |
|
28668
e79e196039a1
fixed and reactivated HOL/Library/Pocklington.thy -- by Mark Hillebrand;
wenzelm
parents:
28228
diff
changeset
|
37 |
Pocklington |
| 15131 | 38 |
Primes |
| 29806 | 39 |
Quickcheck |
| 24615 | 40 |
Quicksort |
| 15131 | 41 |
Quotient |
|
21635
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents:
21256
diff
changeset
|
42 |
Ramsey |
| 29806 | 43 |
Random |
| 29650 | 44 |
Reflection |
|
26192
52617dca8386
new theory of red-black trees, an efficient implementation of finite maps.
krauss
parents:
26173
diff
changeset
|
45 |
RBT |
| 21256 | 46 |
State_Monad |
| 29504 | 47 |
Univ_Poly |
| 15131 | 48 |
While_Combinator |
49 |
Word |
|
|
27475
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
huffman
parents:
27368
diff
changeset
|
50 |
Zorn |
| 15131 | 51 |
begin |
| 10253 | 52 |
end |
53 |
(*>*) |