| author | nipkow | 
| Wed, 25 May 2016 17:40:56 +0200 | |
| changeset 63149 | f5dbab18c404 | 
| parent 63122 | dd651e3f7413 | 
| child 63317 | ca187a9f66da | 
| permissions | -rw-r--r-- | 
| 10253 | 1 | (*<*) | 
| 15131 | 2 | theory Library | 
| 15140 | 3 | imports | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50134diff
changeset | 4 | AList | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: 
16109diff
changeset | 5 | BigO | 
| 29994 
6ca6b6bd6e15
add formalization of a type of integers mod 2 to Library
 huffman parents: 
29993diff
changeset | 6 | Bit | 
| 56942 | 7 | BNF_Axiomatization | 
| 62692 
0701f25fac39
moved 'corec' from ssh://hg@bitbucket.org/jasmin_blanchette/nonprim-corec to Isabelle
 blanchet parents: 
62652diff
changeset | 8 | BNF_Corec | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: 
24281diff
changeset | 9 | Boolean_Algebra | 
| 61766 
507b39df1a57
add formalisation of Bourbaki-Witt fixpoint theorem
 Andreas Lochbihler parents: 
61178diff
changeset | 10 | Bourbaki_Witt_Fixpoint | 
| 21256 | 11 | Char_ord | 
| 58626 | 12 | Code_Test | 
| 29026 | 13 | ContNotDenum | 
| 36648 | 14 | Convex | 
| 62652 
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
 Andreas Lochbihler parents: 
62375diff
changeset | 15 | Complete_Partial_Order2 | 
| 26170 | 16 | Countable | 
| 62373 | 17 | Countable_Complete_Lattices | 
| 55075 | 18 | Countable_Set_Type | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: 
48283diff
changeset | 19 | Debug | 
| 50087 | 20 | Diagonal_Subsequence | 
| 60727 | 21 | Disjoint_Sets | 
| 48283 | 22 | Dlist | 
| 51542 | 23 | Extended | 
| 24 | Extended_Nat | |
| 62375 | 25 | Extended_Nonnegative_Real | 
| 51542 | 26 | Extended_Real | 
| 48028 | 27 | FinFun | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28668diff
changeset | 28 | Float | 
| 29688 | 29 | Formal_Power_Series | 
| 31761 
3585bebe49a8
Added Library/Fraction_Field.thy: The fraction field of any integral
 chaieb parents: 
31379diff
changeset | 30 | Fraction_Field | 
| 53953 | 31 | FSet | 
| 21256 | 32 | FuncSet | 
| 48188 | 33 | Function_Division | 
| 51263 
31e786e0e6a7
turned example into library for comparing growth of functions
 haftmann parents: 
51174diff
changeset | 34 | Function_Growth | 
| 29879 | 35 | Fundamental_Theorem_Algebra | 
| 58196 
1b3fbfb85980
theory about lexicographic ordering on functions
 haftmann parents: 
58110diff
changeset | 36 | Fun_Lexorder | 
| 58197 | 37 | Groups_Big_Fun | 
| 37665 | 38 | 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 | 39 | Infinite_Set | 
| 29993 | 40 | Inner_Product | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50134diff
changeset | 41 | IArray | 
| 35032 
7efe662e41b4
separate library theory for type classes combining lattices with various algebraic structures
 haftmann parents: 
34020diff
changeset | 42 | Lattice_Algebras | 
| 30326 
a01b2de0e3e1
constructive version of Cantor's first diagonalization argument
 haftmann parents: 
30261diff
changeset | 43 | Lattice_Syntax | 
| 57998 
8b7508f848ef
rename Quickcheck_Types to Lattice_Constructions and remove quickcheck setup
 Andreas Lochbihler parents: 
57250diff
changeset | 44 | Lattice_Constructions | 
| 58627 | 45 | Linear_Temporal_Logic_on_Streams | 
| 26173 | 46 | ListVector | 
| 58810 | 47 | Lub_Glb | 
| 29708 | 48 | Mapping | 
| 37790 | 49 | Monad_Syntax | 
| 58199 
5fbe474b5da8
explicit theory with additional, less commonly used list operations
 haftmann parents: 
58197diff
changeset | 50 | More_List | 
| 59813 | 51 | Multiset_Order | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: 
24281diff
changeset | 52 | Numeral_Type | 
| 61178 
0b071f72f330
Omega_Words_Fun: Infinite words as functions from nat.
 lammich <lammich@in.tum.de> parents: 
60727diff
changeset | 53 | Omega_Words_Fun | 
| 15470 | 54 | OptionalSugar | 
| 26232 | 55 | Option_ord | 
| 56020 
f92479477c52
introduced antimono; incseq, decseq are now abbreviations for mono and antimono; renamed Library/Continuity to Library/Order_Continuity; removed up_cont; renamed down_cont to down_continuity and generalized to complete_lattices
 hoelzl parents: 
55159diff
changeset | 56 | Order_Continuity | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: 
48283diff
changeset | 57 | Parallel | 
| 15131 | 58 | Permutation | 
| 44227 
78e033e8ba05
get Library/Permutations.thy compiled and working again
 huffman parents: 
44014diff
changeset | 59 | Permutations | 
| 29987 | 60 | Polynomial | 
| 31060 
75d7c7cc8bdb
added theory for explicit equivalence relation in preorders
 haftmann parents: 
30326diff
changeset | 61 | Preorder | 
| 30019 
a2f19e0a28b2
add theory of products as real vector spaces to Library
 huffman parents: 
30018diff
changeset | 62 | Product_Vector | 
| 60162 | 63 | Quadratic_Discriminant | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35100diff
changeset | 64 | Quotient_List | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35100diff
changeset | 65 | Quotient_Option | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35100diff
changeset | 66 | Quotient_Product | 
| 45074 | 67 | Quotient_Set | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35100diff
changeset | 68 | Quotient_Sum | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35100diff
changeset | 69 | 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 | 70 | Quotient_Type | 
| 21635 
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
 krauss parents: 
21256diff
changeset | 71 | Ramsey | 
| 29650 | 72 | Reflection | 
| 44818 
27ba81ad0890
theory of saturated naturals contributed by Peter Gammie
 haftmann parents: 
44561diff
changeset | 73 | Saturated | 
| 38622 | 74 | Set_Algebras | 
| 63122 | 75 | Set_Permutations | 
| 21256 | 76 | State_Monad | 
| 63071 | 77 | Stirling | 
| 58607 
1f90ea1b4010
move Stream theory from Datatype_Examples to Library
 hoelzl parents: 
58199diff
changeset | 78 | Stream | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
50134diff
changeset | 79 | Sublist | 
| 41474 | 80 | Sum_of_Squares | 
| 33649 
854173fcd21c
added a tabled implementation of the reflexive transitive closure
 bulwahn parents: 
33356diff
changeset | 81 | Transitive_Closure_Table | 
| 59928 | 82 | Tree_Multiset | 
| 15131 | 83 | While_Combinator | 
| 84 | begin | |
| 10253 | 85 | end | 
| 86 | (*>*) |