| author | ballarin | 
| Tue, 16 Dec 2008 14:29:05 +0100 | |
| changeset 29216 | 528e68bea04d | 
| parent 29026 | 5fbaa05f637f | 
| child 29197 | 6d4cb27ed19c | 
| permissions | -rw-r--r-- | 
| 20809 | 1 | (* $Id$ *) | 
| 10253 | 2 | (*<*) | 
| 15131 | 3 | theory Library | 
| 15140 | 4 | imports | 
| 27652 
818666de6c24
refined code generator setup for rational numbers; more simplification rules for rational numbers
 haftmann parents: 
27475diff
changeset | 5 | Abstract_Rat | 
| 21256 | 6 | AssocList | 
| 16908 
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
 avigad parents: 
16109diff
changeset | 7 | BigO | 
| 21256 | 8 | Binomial | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: 
24281diff
changeset | 9 | Boolean_Algebra | 
| 21256 | 10 | Char_ord | 
| 26170 | 11 | Code_Char_chr | 
| 24994 
c385c4eabb3b
consolidated naming conventions for code generator theories
 haftmann parents: 
24626diff
changeset | 12 | Code_Index | 
| 26170 | 13 | Code_Integer | 
| 21256 | 14 | Coinductive_List | 
| 15 | Commutative_Ring | |
| 15131 | 16 | Continuity | 
| 29026 | 17 | ContNotDenum | 
| 26170 | 18 | Countable | 
| 26157 | 19 | Dense_Linear_Order | 
| 23854 | 20 | Efficient_Nat | 
| 26348 | 21 | Enum | 
| 24281 | 22 | Eval_Witness | 
| 23854 | 23 | Executable_Set | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28668diff
changeset | 24 | Float | 
| 21256 | 25 | FuncSet | 
| 26170 | 26 | Imperative_HOL | 
| 27475 
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
 huffman parents: 
27368diff
changeset | 27 | Infinite_Set | 
| 26173 | 28 | ListVector | 
| 15131 | 29 | Multiset | 
| 30 | Nat_Infinity | |
| 31 | Nested_Environment | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: 
24281diff
changeset | 32 | Numeral_Type | 
| 15470 | 33 | OptionalSugar | 
| 26232 | 34 | Option_ord | 
| 15131 | 35 | Permutation | 
| 28668 
e79e196039a1
fixed and reactivated HOL/Library/Pocklington.thy -- by Mark Hillebrand;
 wenzelm parents: 
28228diff
changeset | 36 | Pocklington | 
| 15131 | 37 | Primes | 
| 24615 | 38 | Quicksort | 
| 15131 | 39 | Quotient | 
| 21635 
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
 krauss parents: 
21256diff
changeset | 40 | Ramsey | 
| 26192 
52617dca8386
new theory of red-black trees, an efficient implementation of finite maps.
 krauss parents: 
26173diff
changeset | 41 | RBT | 
| 21256 | 42 | State_Monad | 
| 15131 | 43 | While_Combinator | 
| 44 | Word | |
| 27475 
61b979a2c820
add Infinite_Set and Zorn back in (since they are no longer included in main HOL image)
 huffman parents: 
27368diff
changeset | 45 | Zorn | 
| 15131 | 46 | begin | 
| 10253 | 47 | end | 
| 48 | (*>*) |