author | paulson <lp15@cam.ac.uk> |
Mon, 25 Oct 2021 13:56:08 +0100 | |
changeset 74574 | cff477b6d015 |
parent 74334 | ead56ad40e15 |
child 75801 | 5c1856aaf03d |
permissions | -rw-r--r-- |
73034 | 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:
50134
diff
changeset
|
4 |
AList |
68061
81d90f830f99
clarified theory location and imports: avoid surprises due to Pure instead of Main (e.g. simpset operations);
wenzelm
parents:
67224
diff
changeset
|
5 |
Adhoc_Overloading |
16908
d374530bfaaa
Added two new theories to HOL/Library: SetsAndFunctions.thy and BigO.thy
avigad
parents:
16109
diff
changeset
|
6 |
BigO |
56942 | 7 |
BNF_Axiomatization |
62692
0701f25fac39
moved 'corec' from ssh://hg@bitbucket.org/jasmin_blanchette/nonprim-corec to Isabelle
blanchet
parents:
62652
diff
changeset
|
8 |
BNF_Corec |
61766
507b39df1a57
add formalisation of Bourbaki-Witt fixpoint theorem
Andreas Lochbihler
parents:
61178
diff
changeset
|
9 |
Bourbaki_Witt_Fixpoint |
21256 | 10 |
Char_ord |
73886
93ba8e3fdcdf
move code setup from Cardinality to separate theory
Andreas Lochbihler <mail@andreas-lochbihler.de>
parents:
73622
diff
changeset
|
11 |
Code_Cardinality |
68155 | 12 |
Code_Lazy |
58626 | 13 |
Code_Test |
63377
64adf4ba9526
combinator to build partial equivalence relations from a predicate and an equivalenc relation
haftmann
parents:
63375
diff
changeset
|
14 |
Combine_PER |
62652
7248d106c607
move Complete_Partial_Orders2 from AFP/Coinductive to HOL/Library
Andreas Lochbihler
parents:
62375
diff
changeset
|
15 |
Complete_Partial_Order2 |
67224 | 16 |
Conditional_Parametricity |
71393
fce780f9c9c6
new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
71036
diff
changeset
|
17 |
Confluence |
fce780f9c9c6
new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
71036
diff
changeset
|
18 |
Confluent_Quotient |
26170 | 19 |
Countable |
62373 | 20 |
Countable_Complete_Lattices |
55075 | 21 |
Countable_Set_Type |
48427
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
48283
diff
changeset
|
22 |
Debug |
50087 | 23 |
Diagonal_Subsequence |
66797 | 24 |
Discrete |
60727 | 25 |
Disjoint_Sets |
73326 | 26 |
Disjoint_FSets |
48283 | 27 |
Dlist |
69909
5382f5691a11
proper theory for type of dual ordered lattice in distribution
haftmann
parents:
69790
diff
changeset
|
28 |
Dual_Ordered_Lattice |
69735
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
paulson <lp15@cam.ac.uk>
parents:
69194
diff
changeset
|
29 |
Equipollence |
51542 | 30 |
Extended |
31 |
Extended_Nat |
|
62375 | 32 |
Extended_Nonnegative_Real |
51542 | 33 |
Extended_Real |
63885
a6cd18af8bf9
new type for finite maps; use it in HOL-Probability
Lars Hupel <lars.hupel@mytum.de>
parents:
63762
diff
changeset
|
34 |
Finite_Map |
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28668
diff
changeset
|
35 |
Float |
53953 | 36 |
FSet |
68188
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents:
68155
diff
changeset
|
37 |
FuncSet |
48188 | 38 |
Function_Division |
58196
1b3fbfb85980
theory about lexicographic ordering on functions
haftmann
parents:
58110
diff
changeset
|
39 |
Fun_Lexorder |
66488
9d83e8fe3de3
HOL-Library: going_to filter
Manuel Eberl <eberlm@in.tum.de>
parents:
66451
diff
changeset
|
40 |
Going_To_Filter |
58197 | 41 |
Groups_Big_Fun |
37665 | 42 |
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
|
43 |
Infinite_Set |
71035 | 44 |
Interval |
71036 | 45 |
Interval_Float |
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
50134
diff
changeset
|
46 |
IArray |
68246
b48bab511939
Moved Landau_Symbols from the AFP to HOL-Library
Manuel Eberl <eberlm@in.tum.de>
parents:
68188
diff
changeset
|
47 |
Landau_Symbols |
35032
7efe662e41b4
separate library theory for type classes combining lattices with various algebraic structures
haftmann
parents:
34020
diff
changeset
|
48 |
Lattice_Algebras |
57998
8b7508f848ef
rename Quickcheck_Types to Lattice_Constructions and remove quickcheck setup
Andreas Lochbihler
parents:
57250
diff
changeset
|
49 |
Lattice_Constructions |
58627 | 50 |
Linear_Temporal_Logic_on_Streams |
26173 | 51 |
ListVector |
58810 | 52 |
Lub_Glb |
29708 | 53 |
Mapping |
37790 | 54 |
Monad_Syntax |
58199
5fbe474b5da8
explicit theory with additional, less commonly used list operations
haftmann
parents:
58197
diff
changeset
|
55 |
More_List |
59813 | 56 |
Multiset_Order |
64588 | 57 |
Nonpos_Ints |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
24281
diff
changeset
|
58 |
Numeral_Type |
61178
0b071f72f330
Omega_Words_Fun: Infinite words as functions from nat.
lammich <lammich@in.tum.de>
parents:
60727
diff
changeset
|
59 |
Omega_Words_Fun |
66270
403d84138c5c
State_Monad ~> Open_State_Syntax
Lars Hupel <lars.hupel@mytum.de>
parents:
66015
diff
changeset
|
60 |
Open_State_Syntax |
26232 | 61 |
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:
55159
diff
changeset
|
62 |
Order_Continuity |
48427
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
48283
diff
changeset
|
63 |
Parallel |
66451
5be0b0604d71
syntax for pattern aliases
Lars Hupel <lars.hupel@mytum.de>
parents:
66271
diff
changeset
|
64 |
Pattern_Aliases |
64588 | 65 |
Periodic_Fun |
70042
45787384ff86
new theory Library/Poly_Mapping, of almost-everywhere-zero functions
paulson <lp15@cam.ac.uk>
parents:
69909
diff
changeset
|
66 |
Poly_Mapping |
69790
154cf64e403e
Exponentiation by squaring, fast modular exponentiation
Manuel Eberl <eberlm@in.tum.de>
parents:
69735
diff
changeset
|
67 |
Power_By_Squaring |
31060
75d7c7cc8bdb
added theory for explicit equivalence relation in preorders
haftmann
parents:
30326
diff
changeset
|
68 |
Preorder |
63972 | 69 |
Product_Plus |
60162 | 70 |
Quadratic_Discriminant |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35100
diff
changeset
|
71 |
Quotient_List |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35100
diff
changeset
|
72 |
Quotient_Option |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35100
diff
changeset
|
73 |
Quotient_Product |
45074 | 74 |
Quotient_Set |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35100
diff
changeset
|
75 |
Quotient_Sum |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
35100
diff
changeset
|
76 |
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
|
77 |
Quotient_Type |
21635
32f3e1127de2
added Ramsey.thy to Library imports, to include it in the daily builds
krauss
parents:
21256
diff
changeset
|
78 |
Ramsey |
29650 | 79 |
Reflection |
64588 | 80 |
Rewrite |
44818
27ba81ad0890
theory of saturated naturals contributed by Peter Gammie
haftmann
parents:
44561
diff
changeset
|
81 |
Saturated |
38622 | 82 |
Set_Algebras |
69004
f6a0c8115e9c
Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
68246
diff
changeset
|
83 |
Set_Idioms |
72281
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
71956
diff
changeset
|
84 |
Signed_Division |
66271 | 85 |
State_Monad |
58607
1f90ea1b4010
move Stream theory from Datatype_Examples to Library
hoelzl
parents:
58199
diff
changeset
|
86 |
Stream |
69194 | 87 |
Sorting_Algorithms |
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
50134
diff
changeset
|
88 |
Sublist |
41474 | 89 |
Sum_of_Squares |
33649
854173fcd21c
added a tabled implementation of the reflexive transitive closure
bulwahn
parents:
33356
diff
changeset
|
90 |
Transitive_Closure_Table |
59928 | 91 |
Tree_Multiset |
66510 | 92 |
Tree_Real |
63762 | 93 |
Type_Length |
66563 | 94 |
Uprod |
15131 | 95 |
While_Combinator |
72515
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
haftmann
parents:
72281
diff
changeset
|
96 |
Word |
70342
e4d626692640
clear separation of types for bits (False / True) and Z2 (0 / 1)
haftmann
parents:
70042
diff
changeset
|
97 |
Z2 |
15131 | 98 |
begin |
10253 | 99 |
end |
66451
5be0b0604d71
syntax for pattern aliases
Lars Hupel <lars.hupel@mytum.de>
parents:
66271
diff
changeset
|
100 |
(*>*) |