src/HOL/Library/Library.thy
author haftmann
Sat, 24 Dec 2011 15:53:10 +0100
changeset 45970 b6d0cff57d96
parent 45748 cf79cc09cab4
child 45990 b7b905b23b2a
permissions -rw-r--r--
adjusted to set/pred distinction by means of type constructor `set`
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10253
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
     1
(*<*)
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
     2
theory Library
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
     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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    10
  Char_ord
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    11
  Continuity
29026
5fbaa05f637f moved ContNotDenum into Library
nipkow
parents: 28952
diff changeset
    12
  ContNotDenum
36648
43b66dcd9266 Add Convex to Library build
hoelzl
parents: 36147
diff changeset
    13
  Convex
26170
66e6b967ccf1 added theories for imperative HOL
haftmann
parents: 26157
diff changeset
    14
  Countable
43976
af17d7934116 Add theory for setting up monad syntax for Cset
Andreas Lochbihler
parents: 43958
diff changeset
    15
  Cset_Monad
43146
09f74fda1b1d splitting Dlist theory in Dlist and Dlist_Cset
bulwahn
parents: 43124
diff changeset
    16
  Dlist_Cset
24281
7d0334b69711 added Eval_Witness theory
haftmann
parents: 24197
diff changeset
    17
  Eval_Witness
43919
a7e4fb1a0502 rename Nat_Infinity (inat) to Extended_Nat (enat)
hoelzl
parents: 43241
diff changeset
    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
6ed9ac8410d8 Added Formal_Power_Series in imports
chaieb
parents: 29650
diff changeset
    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
6b1ccda8bf19 move FrechetDeriv.thy to Library
huffman
parents: 29985
diff changeset
    22
  FrechetDeriv
40672
abd4e7358847 replaced misleading Fset/fset name -- these do not stand for finite sets
haftmann
parents: 40349
diff changeset
    23
  Cset
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    24
  FuncSet
38622
86fc906dcd86 split and enriched theory SetsAndFunctions
haftmann
parents: 37818
diff changeset
    25
  Function_Algebras
29879
4425849f5db7 Moved FTA into Lib and cleaned it up a little.
nipkow
parents: 29847
diff changeset
    26
  Fundamental_Theorem_Algebra
37665
579258a77fec Add theory for indicator function.
hoelzl
parents: 37023
diff changeset
    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
84b2c432b94a new theory of real inner product spaces
huffman
parents: 29987
diff changeset
    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
5cac519abe4e renamed ListSpace to ListVector;
wenzelm
parents: 26170
diff changeset
    32
  ListVector
31990
1d4d0b305f16 move Kleene_Algebra to Library
krauss
parents: 31849
diff changeset
    33
  Kleene_Algebra
29708
e40b70d38909 added Mapping.thy to Library
haftmann
parents: 29688
diff changeset
    34
  Mapping
37790
7fea92005066 uniform do notation for monads
krauss
parents: 37789
diff changeset
    35
  Monad_Syntax
37023
efc202e1677e added theory More_List
haftmann
parents: 36962
diff changeset
    36
  More_List
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    37
  Multiset
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents: 24281
diff changeset
    38
  Numeral_Type
44013
5cfc1c36ae97 moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents: 43976
diff changeset
    39
  Old_Recdef
15470
7e12ad2f6672 added OptionalSugar
nipkow
parents: 15324
diff changeset
    40
  OptionalSugar
26232
075264a0a4bc canonical order on option type
haftmann
parents: 26192
diff changeset
    41
  Option_ord
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    42
  Permutation
44227
78e033e8ba05 get Library/Permutations.thy compiled and working again
huffman
parents: 44014
diff changeset
    43
  Permutations
29985
57975b45ab70 split polynomial-related stuff from Deriv.thy into Library/Poly_Deriv.thy
huffman
parents: 29879
diff changeset
    44
  Poly_Deriv
29987
391dcbd7e4dd move Polynomial.thy to Library
huffman
parents: 29986
diff changeset
    45
  Polynomial
31060
75d7c7cc8bdb added theory for explicit equivalence relation in preorders
haftmann
parents: 30326
diff changeset
    46
  Preorder
30019
a2f19e0a28b2 add theory of products as real vector spaces to Library
huffman
parents: 30018
diff changeset
    47
  Product_Vector
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35100
diff changeset
    48
  Quotient_List
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35100
diff changeset
    49
  Quotient_Option
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35100
diff changeset
    50
  Quotient_Product
45074
04286b0fc856 Quotient_Set.thy is part of library
haftmann
parents: 44897
diff changeset
    51
  Quotient_Set
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35100
diff changeset
    52
  Quotient_Sum
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35100
diff changeset
    53
  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
    54
  Quotient_Type
21635
32f3e1127de2 added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents: 21256
diff changeset
    55
  Ramsey
29650
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 29504
diff changeset
    56
  Reflection
43124
fdb7e1d5f762 splitting RBT theory into RBT and RBT_Mapping
bulwahn
parents: 41474
diff changeset
    57
  RBT_Mapping
44818
27ba81ad0890 theory of saturated naturals contributed by Peter Gammie
haftmann
parents: 44561
diff changeset
    58
  Saturated
38622
86fc906dcd86 split and enriched theory SetsAndFunctions
haftmann
parents: 37818
diff changeset
    59
  Set_Algebras
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    60
  State_Monad
41474
60d091240485 renamed Sum_Of_Squares to Sum_of_Squares;
wenzelm
parents: 40673
diff changeset
    61
  Sum_of_Squares
33649
854173fcd21c added a tabled implementation of the reflexive transitive closure
bulwahn
parents: 33356
diff changeset
    62
  Transitive_Closure_Table
29504
4c3441f2f619 moved Univ_Poly to Library
haftmann
parents: 29399
diff changeset
    63
  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
    64
  Wfrec
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    65
  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
    66
  Zorn
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    67
begin
10253
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
    68
end
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
    69
(*>*)