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