src/HOL/Library/Library.thy
author kleing
Mon, 20 Aug 2007 00:22:18 +0200
changeset 24332 e3a2b75b1cf9
parent 24281 7d0334b69711
child 24530 1bac25879117
permissions -rw-r--r--
boolean algebras as locales and numbers as types by Brian Huffman
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20809
6c4fd0b4b63a moved theory Infinite_Set to Library;
wenzelm
parents: 20400
diff changeset
     1
(* $Id$ *)
10253
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
     2
(*<*)
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
     3
theory Library
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     4
imports
24197
c9e3cb5e5681 proper implementation of rational numbers
haftmann
parents: 23854
diff changeset
     5
  Abstract_Rat
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
     6
  AssocList
16908
d374530bfaaa Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents: 16109
diff changeset
     7
  BigO
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
     8
  Binomial
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
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    11
  Coinductive_List
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    12
  Commutative_Ring
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    13
  Continuity
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents: 23192
diff changeset
    14
  Efficient_Nat
22519
eb70ed79dac7 importing Eval theory
haftmann
parents: 22359
diff changeset
    15
  Eval
24281
7d0334b69711 added Eval_Witness theory
haftmann
parents: 24197
diff changeset
    16
  Eval_Witness
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents: 23192
diff changeset
    17
  Executable_Rat
22981
cf071f3fc4ae A verified theory for rational numbers representation and simple calculations;
chaieb
parents: 22799
diff changeset
    18
  Executable_Real
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents: 23192
diff changeset
    19
  Executable_Set
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    20
  FuncSet
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    21
  GCD
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    22
  Infinite_Set
23854
688a8a7bcd4e uniform naming conventions for CG theories
haftmann
parents: 23192
diff changeset
    23
  ML_String
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    24
  Multiset
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    25
  NatPair
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    26
  Nat_Infinity
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    27
  Nested_Environment
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents: 24281
diff changeset
    28
  Numeral_Type
15470
7e12ad2f6672 added OptionalSugar
nipkow
parents: 15324
diff changeset
    29
  OptionalSugar
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    30
  Parity
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    31
  Permutation
22799
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents: 22519
diff changeset
    32
  Pretty_Char_chr
ed7d53db2170 moved code generation pretty integers and characters to separate theories
haftmann
parents: 22519
diff changeset
    33
  Pretty_Int
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    34
  Primes
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    35
  Quotient
21635
32f3e1127de2 added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents: 21256
diff changeset
    36
  Ramsey
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 21192
diff changeset
    37
  State_Monad
22359
94a794672c8b Added formalization of size-change principle (experimental).
krauss
parents: 21635
diff changeset
    38
  Size_Change_Termination
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    39
  While_Combinator
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    40
  Word
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    41
  Zorn
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    42
begin
10253
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
    43
end
73b46b18c348 "The Supplemental Isabelle/HOL Library";
wenzelm
parents:
diff changeset
    44
(*>*)