| author | wenzelm | 
| Mon, 26 Mar 2012 21:03:30 +0200 | |
| changeset 47133 | 89b13238d7f2 | 
| parent 45990 | b7b905b23b2a | 
| child 47232 | e2f0176149d0 | 
| 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  | 
| 
44897
 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 
bulwahn 
parents: 
44818 
diff
changeset
 | 
5  | 
AList_Mapping  | 
| 
16908
 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 
avigad 
parents: 
16109 
diff
changeset
 | 
6  | 
BigO  | 
| 21256 | 7  | 
Binomial  | 
| 
29994
 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 
huffman 
parents: 
29993 
diff
changeset
 | 
8  | 
Bit  | 
| 
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  | 
| 15131 | 11  | 
Continuity  | 
| 29026 | 12  | 
ContNotDenum  | 
| 36648 | 13  | 
Convex  | 
| 26170 | 14  | 
Countable  | 
| 
43976
 
af17d7934116
Add theory for setting up monad syntax for Cset
 
Andreas Lochbihler 
parents: 
43958 
diff
changeset
 | 
15  | 
Cset_Monad  | 
| 43146 | 16  | 
Dlist_Cset  | 
| 24281 | 17  | 
Eval_Witness  | 
| 43919 | 18  | 
Extended_Nat  | 
| 
28952
 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 
haftmann 
parents: 
28668 
diff
changeset
 | 
19  | 
Float  | 
| 29688 | 20  | 
Formal_Power_Series  | 
| 
31761
 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 
chaieb 
parents: 
31379 
diff
changeset
 | 
21  | 
Fraction_Field  | 
| 29986 | 22  | 
FrechetDeriv  | 
| 
40672
 
abd4e7358847
replaced misleading Fset/fset name -- these do not stand for finite sets
 
haftmann 
parents: 
40349 
diff
changeset
 | 
23  | 
Cset  | 
| 21256 | 24  | 
FuncSet  | 
| 38622 | 25  | 
Function_Algebras  | 
| 29879 | 26  | 
Fundamental_Theorem_Algebra  | 
| 37665 | 27  | 
Indicator_Function  | 
| 
27475
 
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
 
huffman 
parents: 
27368 
diff
changeset
 | 
28  | 
Infinite_Set  | 
| 29993 | 29  | 
Inner_Product  | 
| 
35032
 
7efe662e41b4
separate library theory for type classes combining lattices with various algebraic structures
 
haftmann 
parents: 
34020 
diff
changeset
 | 
30  | 
Lattice_Algebras  | 
| 
30326
 
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
 
haftmann 
parents: 
30261 
diff
changeset
 | 
31  | 
Lattice_Syntax  | 
| 26173 | 32  | 
ListVector  | 
| 31990 | 33  | 
Kleene_Algebra  | 
| 29708 | 34  | 
Mapping  | 
| 37790 | 35  | 
Monad_Syntax  | 
| 15131 | 36  | 
Multiset  | 
| 
24332
 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 
kleing 
parents: 
24281 
diff
changeset
 | 
37  | 
Numeral_Type  | 
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
43976 
diff
changeset
 | 
38  | 
Old_Recdef  | 
| 15470 | 39  | 
OptionalSugar  | 
| 26232 | 40  | 
Option_ord  | 
| 15131 | 41  | 
Permutation  | 
| 
44227
 
78e033e8ba05
get Library/Permutations.thy compiled and working again
 
huffman 
parents: 
44014 
diff
changeset
 | 
42  | 
Permutations  | 
| 
29985
 
57975b45ab70
split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
 
huffman 
parents: 
29879 
diff
changeset
 | 
43  | 
Poly_Deriv  | 
| 29987 | 44  | 
Polynomial  | 
| 
31060
 
75d7c7cc8bdb
added theory for explicit equivalence relation in preorders
 
haftmann 
parents: 
30326 
diff
changeset
 | 
45  | 
Preorder  | 
| 
30019
 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 
huffman 
parents: 
30018 
diff
changeset
 | 
46  | 
Product_Vector  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35100 
diff
changeset
 | 
47  | 
Quotient_List  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35100 
diff
changeset
 | 
48  | 
Quotient_Option  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35100 
diff
changeset
 | 
49  | 
Quotient_Product  | 
| 45074 | 50  | 
Quotient_Set  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35100 
diff
changeset
 | 
51  | 
Quotient_Sum  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35100 
diff
changeset
 | 
52  | 
Quotient_Syntax  | 
| 
35100
 
53754ec7360b
renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid clash with new theory Quotient in Main HOL;
 
wenzelm 
parents: 
35091 
diff
changeset
 | 
53  | 
Quotient_Type  | 
| 
21635
 
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
 
krauss 
parents: 
21256 
diff
changeset
 | 
54  | 
Ramsey  | 
| 29650 | 55  | 
Reflection  | 
| 43124 | 56  | 
RBT_Mapping  | 
| 
44818
 
27ba81ad0890
theory of saturated naturals contributed by Peter Gammie
 
haftmann 
parents: 
44561 
diff
changeset
 | 
57  | 
Saturated  | 
| 38622 | 58  | 
Set_Algebras  | 
| 21256 | 59  | 
State_Monad  | 
| 41474 | 60  | 
Sum_of_Squares  | 
| 
33649
 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 
bulwahn 
parents: 
33356 
diff
changeset
 | 
61  | 
Transitive_Closure_Table  | 
| 29504 | 62  | 
Univ_Poly  | 
| 
44014
 
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
 
krauss 
parents: 
44013 
diff
changeset
 | 
63  | 
Wfrec  | 
| 15131 | 64  | 
While_Combinator  | 
| 
27475
 
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
 
huffman 
parents: 
27368 
diff
changeset
 | 
65  | 
Zorn  | 
| 15131 | 66  | 
begin  | 
| 10253 | 67  | 
end  | 
68  | 
(*>*)  |