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