| author | blanchet | 
| Sat, 08 Sep 2012 21:30:31 +0200 | |
| changeset 49222 | cbe8c859817c | 
| parent 49110 | 2e43fb45b91b | 
| child 49310 | 6e30078de4f0 | 
| permissions | -rw-r--r-- | 
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 1 | session HOL (main) = Pure + | 
| 48338 | 2 |   description {* Classical Higher-order Logic *}
 | 
| 3 | options [document_graph] | |
| 4 | theories Complex_Main | |
| 48901 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 5 | files | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 6 | "Tools/Quickcheck/Narrowing_Engine.hs" | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 7 | "Tools/Quickcheck/PNF_Narrowing_Engine.hs" | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 8 | "document/root.bib" | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 9 | "document/root.tex" | 
| 48338 | 10 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 11 | session "HOL-Base" = Pure + | 
| 48338 | 12 |   description {* Raw HOL base, with minimal tools *}
 | 
| 48470 
7483aa690b4f
clarified "document" again, eliminated redundant "no_document";
 wenzelm parents: 
48458diff
changeset | 13 | options [document = false] | 
| 48338 | 14 | theories HOL | 
| 15 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 16 | session "HOL-Plain" = Pure + | 
| 48338 | 17 |   description {* HOL side-entry after bootstrap of many tools and packages *}
 | 
| 48470 
7483aa690b4f
clarified "document" again, eliminated redundant "no_document";
 wenzelm parents: 
48458diff
changeset | 18 | options [document = false] | 
| 48338 | 19 | theories Plain | 
| 20 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 21 | session "HOL-Main" = Pure + | 
| 48338 | 22 |   description {* HOL side-entry for Main only, without Complex_Main *}
 | 
| 48470 
7483aa690b4f
clarified "document" again, eliminated redundant "no_document";
 wenzelm parents: 
48458diff
changeset | 23 | options [document = false] | 
| 48338 | 24 | theories Main | 
| 48901 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 25 | files | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 26 | "Tools/Quickcheck/Narrowing_Engine.hs" | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 27 | "Tools/Quickcheck/PNF_Narrowing_Engine.hs" | 
| 48338 | 28 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 29 | session "HOL-Proofs" = Pure + | 
| 48509 | 30 |   description {* HOL-Main with explicit proof terms *}
 | 
| 48470 
7483aa690b4f
clarified "document" again, eliminated redundant "no_document";
 wenzelm parents: 
48458diff
changeset | 31 | options [document = false, proofs = 2, parallel_proofs = 0] | 
| 48338 | 32 | theories Main | 
| 48901 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 33 | files | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 34 | "Tools/Quickcheck/Narrowing_Engine.hs" | 
| 
5e0455e29339
more basic file dependencies -- no load command here;
 wenzelm parents: 
48765diff
changeset | 35 | "Tools/Quickcheck/PNF_Narrowing_Engine.hs" | 
| 48338 | 36 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 37 | session "HOL-Library" in Library = HOL + | 
| 48481 | 38 |   description {* Classical Higher-order Logic -- batteries included *}
 | 
| 39 | theories | |
| 40 | Library | |
| 49077 
154f25a162e3
renamed theory List_Prefix into Sublist (since it is not only about prefixes)
 Christian Sternagel parents: 
48984diff
changeset | 41 | Sublist | 
| 48481 | 42 | List_lexord | 
| 43 | Sublist_Order | |
| 44 | Product_Lattice | |
| 45 | Code_Char_chr | |
| 46 | Code_Char_ord | |
| 47 | Code_Integer | |
| 48 | Efficient_Nat | |
| 48721 | 49 | (* Code_Prolog FIXME cf. 76965c356d2a *) | 
| 48481 | 50 | Code_Real_Approx_By_Float | 
| 51 | Target_Numeral | |
| 48932 
c6e679443adc
actual use of (sos remote_csdp) via ISABELLE_FULL_TEST;
 wenzelm parents: 
48901diff
changeset | 52 | theories [condition = ISABELLE_FULL_TEST] | 
| 
c6e679443adc
actual use of (sos remote_csdp) via ISABELLE_FULL_TEST;
 wenzelm parents: 
48901diff
changeset | 53 | Sum_of_Squares_Remote | 
| 48481 | 54 | files "document/root.bib" "document/root.tex" | 
| 55 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 56 | session "HOL-Hahn_Banach" in Hahn_Banach = HOL + | 
| 48481 | 57 |   description {*
 | 
| 58 | Author: Gertrud Bauer, TU Munich | |
| 59 | ||
| 60 | The Hahn-Banach theorem for real vector spaces. | |
| 61 | *} | |
| 62 | options [document_graph] | |
| 63 | theories Hahn_Banach | |
| 64 | files "document/root.bib" "document/root.tex" | |
| 65 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 66 | session "HOL-Induct" in Induct = HOL + | 
| 48481 | 67 | theories [quick_and_dirty] | 
| 68 | Common_Patterns | |
| 69 | theories | |
| 70 | QuoDataType | |
| 71 | QuoNestedDataType | |
| 72 | Term | |
| 73 | SList | |
| 74 | ABexp | |
| 75 | Tree | |
| 76 | Ordinals | |
| 77 | Sigma_Algebra | |
| 78 | Comb | |
| 79 | PropLog | |
| 80 | Com | |
| 81 | files "document/root.tex" | |
| 82 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 83 | session "HOL-IMP" in IMP = HOL + | 
| 48481 | 84 | options [document_graph] | 
| 85 | theories [document = false] | |
| 86 | "~~/src/HOL/ex/Interpretation_with_Defs" | |
| 87 | "~~/src/HOL/Library/While_Combinator" | |
| 88 | "~~/src/HOL/Library/Char_ord" | |
| 89 | "~~/src/HOL/Library/List_lexord" | |
| 90 | theories | |
| 91 | BExp | |
| 92 | ASM | |
| 93 | Small_Step | |
| 94 | Denotation | |
| 95 | Comp_Rev | |
| 96 | Poly_Types | |
| 97 | Sec_Typing | |
| 98 | Sec_TypingT | |
| 99 | Def_Ass_Sound_Big | |
| 100 | Def_Ass_Sound_Small | |
| 101 | Live | |
| 102 | Live_True | |
| 103 | Hoare_Examples | |
| 104 | VC | |
| 105 | HoareT | |
| 106 | Collecting1 | |
| 48765 
fb1ed5230abc
special code with lists no longer necessary, use sets
 nipkow parents: 
48738diff
changeset | 107 | Collecting_Examples | 
| 48481 | 108 | Abs_Int_Tests | 
| 109 | Abs_Int1_parity | |
| 110 | Abs_Int1_const | |
| 111 | Abs_Int3 | |
| 112 | "Abs_Int_ITP/Abs_Int1_parity_ITP" | |
| 113 | "Abs_Int_ITP/Abs_Int1_const_ITP" | |
| 114 | "Abs_Int_ITP/Abs_Int3_ITP" | |
| 115 | "Abs_Int_Den/Abs_Int_den2" | |
| 116 | Procs_Dyn_Vars_Dyn | |
| 117 | Procs_Stat_Vars_Dyn | |
| 118 | Procs_Stat_Vars_Stat | |
| 119 | C_like | |
| 120 | OO | |
| 121 | Fold | |
| 122 | files "document/root.bib" "document/root.tex" | |
| 123 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 124 | session "HOL-IMPP" in IMPP = HOL + | 
| 48481 | 125 |   description {*
 | 
| 126 | Author: David von Oheimb | |
| 127 | Copyright 1999 TUM | |
| 128 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 129 | options [document = false] | 
| 48481 | 130 | theories EvenOdd | 
| 131 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 132 | session "HOL-Import" in Import = HOL + | 
| 48481 | 133 | options [document_graph] | 
| 134 | theories HOL_Light_Maps | |
| 135 | theories [condition = HOL_LIGHT_BUNDLE] HOL_Light_Import | |
| 136 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 137 | session "HOL-Number_Theory" in Number_Theory = HOL + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 138 | options [document = false] | 
| 48481 | 139 | theories Number_Theory | 
| 140 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 141 | session "HOL-Old_Number_Theory" in Old_Number_Theory = HOL + | 
| 48481 | 142 | options [document_graph] | 
| 143 | theories [document = false] | |
| 144 | "~~/src/HOL/Library/Infinite_Set" | |
| 145 | "~~/src/HOL/Library/Permutation" | |
| 146 | theories | |
| 147 | Fib | |
| 148 | Factorization | |
| 149 | Chinese | |
| 150 | WilsonRuss | |
| 151 | WilsonBij | |
| 152 | Quadratic_Reciprocity | |
| 153 | Primes | |
| 154 | Pocklington | |
| 155 | files "document/root.tex" | |
| 156 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 157 | session "HOL-Hoare" in Hoare = HOL + | 
| 48481 | 158 | theories Hoare | 
| 159 | files "document/root.bib" "document/root.tex" | |
| 160 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 161 | session "HOL-Hoare_Parallel" in Hoare_Parallel = HOL + | 
| 48481 | 162 | options [document_graph] | 
| 163 | theories Hoare_Parallel | |
| 164 | files "document/root.bib" "document/root.tex" | |
| 165 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 166 | session "HOL-Codegenerator_Test" in Codegenerator_Test = "HOL-Library" + | 
| 48481 | 167 | options [document = false, document_graph = false, browser_info = false] | 
| 48624 | 168 | theories Generate Generate_Pretty RBT_Set_Test | 
| 48481 | 169 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 170 | session "HOL-Metis_Examples" in Metis_Examples = HOL + | 
| 48481 | 171 |   description {*
 | 
| 172 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 173 | Author: Jasmin Blanchette, TU Muenchen | |
| 174 | ||
| 175 | Testing Metis and Sledgehammer. | |
| 176 | *} | |
| 48679 | 177 | options [timeout = 3600, document = false] | 
| 48481 | 178 | theories | 
| 179 | Abstraction | |
| 180 | Big_O | |
| 181 | Binary_Tree | |
| 182 | Clausification | |
| 183 | Message | |
| 184 | Proxies | |
| 185 | Tarski | |
| 186 | Trans_Closure | |
| 187 | Sets | |
| 188 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 189 | session "HOL-Nitpick_Examples" in Nitpick_Examples = HOL + | 
| 48481 | 190 |   description {*
 | 
| 191 | Author: Jasmin Blanchette, TU Muenchen | |
| 192 | Copyright 2009 | |
| 193 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 194 | options [document = false] | 
| 48481 | 195 | theories [quick_and_dirty] Nitpick_Examples | 
| 196 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 197 | session "HOL-Algebra" in Algebra = HOL + | 
| 48481 | 198 |   description {*
 | 
| 199 | Author: Clemens Ballarin, started 24 September 1999 | |
| 200 | ||
| 201 | The Isabelle Algebraic Library. | |
| 202 | *} | |
| 203 | options [document_graph] | |
| 204 | theories [document = false] | |
| 205 | (* Preliminaries from set and number theory *) | |
| 206 | "~~/src/HOL/Library/FuncSet" | |
| 207 | "~~/src/HOL/Old_Number_Theory/Primes" | |
| 208 | "~~/src/HOL/Library/Binomial" | |
| 209 | "~~/src/HOL/Library/Permutation" | |
| 210 | theories | |
| 211 | (*** New development, based on explicit structures ***) | |
| 212 | (* Groups *) | |
| 213 | FiniteProduct (* Product operator for commutative groups *) | |
| 214 | Sylow (* Sylow's theorem *) | |
| 215 | Bij (* Automorphism Groups *) | |
| 216 | ||
| 217 | (* Rings *) | |
| 218 | Divisibility (* Rings *) | |
| 219 | IntRing (* Ideals and residue classes *) | |
| 220 | UnivPoly (* Polynomials *) | |
| 221 | theories [document = false] | |
| 222 | (*** Old development, based on axiomatic type classes ***) | |
| 223 | "abstract/Abstract" (*The ring theory*) | |
| 224 | "poly/Polynomial" (*The full theory*) | |
| 225 | files "document/root.bib" "document/root.tex" | |
| 226 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 227 | session "HOL-Auth" in Auth = HOL + | 
| 48481 | 228 | options [document_graph] | 
| 229 | theories | |
| 230 | Auth_Shared | |
| 231 | Auth_Public | |
| 232 | "Smartcard/Auth_Smartcard" | |
| 233 | "Guard/Auth_Guard_Shared" | |
| 234 | "Guard/Auth_Guard_Public" | |
| 235 | files "document/root.tex" | |
| 236 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 237 | session "HOL-UNITY" in UNITY = HOL + | 
| 48481 | 238 |   description {*
 | 
| 239 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 240 | Copyright 1998 University of Cambridge | |
| 241 | ||
| 242 | Verifying security protocols using UNITY. | |
| 243 | *} | |
| 244 | options [document_graph] | |
| 245 | theories [document = false] "../Auth/Public" | |
| 246 | theories | |
| 247 | (*Basic meta-theory*) | |
| 248 | "UNITY_Main" | |
| 249 | ||
| 250 | (*Simple examples: no composition*) | |
| 251 | "Simple/Deadlock" | |
| 252 | "Simple/Common" | |
| 253 | "Simple/Network" | |
| 254 | "Simple/Token" | |
| 255 | "Simple/Channel" | |
| 256 | "Simple/Lift" | |
| 257 | "Simple/Mutex" | |
| 258 | "Simple/Reach" | |
| 259 | "Simple/Reachability" | |
| 260 | ||
| 261 | (*Verifying security protocols using UNITY*) | |
| 262 | "Simple/NSP_Bad" | |
| 263 | ||
| 264 | (*Example of composition*) | |
| 265 | "Comp/Handshake" | |
| 266 | ||
| 267 | (*Universal properties examples*) | |
| 268 | "Comp/Counter" | |
| 269 | "Comp/Counterc" | |
| 270 | "Comp/Priority" | |
| 271 | ||
| 272 | "Comp/TimerArray" | |
| 273 | "Comp/Progress" | |
| 274 | ||
| 275 | "Comp/Alloc" | |
| 276 | "Comp/AllocImpl" | |
| 277 | "Comp/Client" | |
| 278 | ||
| 279 | (*obsolete*) | |
| 280 | "ELT" | |
| 281 | files "document/root.tex" | |
| 282 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 283 | session "HOL-Unix" in Unix = HOL + | 
| 48481 | 284 | options [print_mode = "no_brackets,no_type_brackets"] | 
| 285 | theories Unix | |
| 286 | files "document/root.bib" "document/root.tex" | |
| 287 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 288 | session "HOL-ZF" in ZF = HOL + | 
| 48481 | 289 |   description {* *}
 | 
| 290 | theories MainZF Games | |
| 291 | files "document/root.tex" | |
| 292 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 293 | session "HOL-Imperative_HOL" in Imperative_HOL = HOL + | 
| 48481 | 294 |   description {* *}
 | 
| 295 | options [document_graph, print_mode = "iff,no_brackets"] | |
| 296 | theories [document = false] | |
| 297 | "~~/src/HOL/Library/Countable" | |
| 298 | "~~/src/HOL/Library/Monad_Syntax" | |
| 299 | "~~/src/HOL/Library/Code_Natural" | |
| 300 | "~~/src/HOL/Library/LaTeXsugar" | |
| 301 | theories Imperative_HOL_ex | |
| 302 | files "document/root.bib" "document/root.tex" | |
| 303 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 304 | session "HOL-Decision_Procs" in Decision_Procs = HOL + | 
| 48496 
a7eed34cf219
added condition = ISABELLE_POLYML according to no-smlnj targets in IsaMakefile;
 wenzelm parents: 
48493diff
changeset | 305 | options [condition = ISABELLE_POLYML, document = false] | 
| 48481 | 306 | theories Decision_Procs | 
| 307 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 308 | session "HOL-Proofs-ex" in "Proofs/ex" = "HOL-Proofs" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 309 | options [document = false, proofs = 2, parallel_proofs = 0] | 
| 48481 | 310 | theories Hilbert_Classical | 
| 311 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 312 | session "HOL-Proofs-Extraction" in "Proofs/Extraction" = "HOL-Proofs" + | 
| 48481 | 313 |   description {* Examples for program extraction in Higher-Order Logic *}
 | 
| 48496 
a7eed34cf219
added condition = ISABELLE_POLYML according to no-smlnj targets in IsaMakefile;
 wenzelm parents: 
48493diff
changeset | 314 | options [condition = ISABELLE_POLYML, proofs = 2, parallel_proofs = 0] | 
| 48481 | 315 | theories [document = false] | 
| 316 | "~~/src/HOL/Library/Efficient_Nat" | |
| 317 | "~~/src/HOL/Library/Monad_Syntax" | |
| 318 | "~~/src/HOL/Number_Theory/Primes" | |
| 319 | "~~/src/HOL/Number_Theory/UniqueFactorization" | |
| 320 | "~~/src/HOL/Library/State_Monad" | |
| 321 | theories | |
| 322 | Greatest_Common_Divisor | |
| 323 | Warshall | |
| 324 | Higman_Extraction | |
| 325 | Pigeonhole | |
| 326 | Euclid | |
| 327 | files "document/root.bib" "document/root.tex" | |
| 328 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 329 | session "HOL-Proofs-Lambda" in "Proofs/Lambda" = "HOL-Proofs" + | 
| 48481 | 330 | options [document_graph, print_mode = "no_brackets", proofs = 2, parallel_proofs = 0] | 
| 331 | theories [document = false] | |
| 332 | "~~/src/HOL/Library/Code_Integer" | |
| 333 | theories | |
| 334 | Eta | |
| 335 | StrongNorm | |
| 336 | Standardization | |
| 337 | WeakNorm | |
| 338 | files "document/root.bib" "document/root.tex" | |
| 339 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 340 | session "HOL-Prolog" in Prolog = HOL + | 
| 48481 | 341 |   description {*
 | 
| 342 | Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) | |
| 343 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 344 | options [document = false] | 
| 48481 | 345 | theories Test Type | 
| 346 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 347 | session "HOL-MicroJava" in MicroJava = HOL + | 
| 48481 | 348 | options [document_graph] | 
| 349 | theories [document = false] "~~/src/HOL/Library/While_Combinator" | |
| 350 | theories MicroJava | |
| 351 | files | |
| 352 | "document/introduction.tex" | |
| 353 | "document/root.bib" | |
| 354 | "document/root.tex" | |
| 355 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 356 | session "HOL-MicroJava-skip_proofs" in MicroJava = HOL + | 
| 48636 | 357 | options [condition = ISABELLE_FULL_TEST, document = false, skip_proofs, quick_and_dirty] | 
| 358 | theories MicroJava | |
| 359 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 360 | session "HOL-NanoJava" in NanoJava = HOL + | 
| 48481 | 361 | options [document_graph] | 
| 362 | theories Example | |
| 363 | files "document/root.bib" "document/root.tex" | |
| 364 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 365 | session "HOL-Bali" in Bali = HOL + | 
| 48481 | 366 | options [document_graph] | 
| 367 | theories | |
| 368 | AxExample | |
| 369 | AxSound | |
| 370 | AxCompl | |
| 371 | Trans | |
| 372 | files "document/root.tex" | |
| 373 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 374 | session "HOL-IOA" in IOA = HOL + | 
| 48481 | 375 |   description {*
 | 
| 376 | Author: Tobias Nipkow & Konrad Slind | |
| 377 | Copyright 1994 TU Muenchen | |
| 378 | ||
| 379 | The meta theory of I/O-Automata. | |
| 380 | ||
| 381 |     @inproceedings{Nipkow-Slind-IOA,
 | |
| 382 |     author={Tobias Nipkow and Konrad Slind},
 | |
| 383 |     title={{I/O} Automata in {Isabelle/HOL}},
 | |
| 384 |     booktitle={Proc.\ TYPES Workshop 1994},
 | |
| 385 | publisher=Springer, | |
| 386 | series=LNCS, | |
| 387 |     note={To appear}}
 | |
| 388 | ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/ioa.ps.gz | |
| 389 | ||
| 390 | and | |
| 391 | ||
| 392 |     @inproceedings{Mueller-Nipkow,
 | |
| 393 |     author={Olaf M\"uller and Tobias Nipkow},
 | |
| 394 |     title={Combining Model Checking and Deduction for {I/O}-Automata},
 | |
| 395 |     booktitle={Proc.\ TACAS Workshop},
 | |
| 396 |     organization={Aarhus University, BRICS report},
 | |
| 397 | year=1995} | |
| 398 | ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/tacas.dvi.gz | |
| 399 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 400 | options [document = false] | 
| 48481 | 401 | theories Solve | 
| 402 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 403 | session "HOL-Lattice" in Lattice = HOL + | 
| 48481 | 404 |   description {*
 | 
| 405 | Author: Markus Wenzel, TU Muenchen | |
| 406 | ||
| 407 | Basic theory of lattices and orders. | |
| 408 | *} | |
| 409 | theories CompleteLattice | |
| 410 | files "document/root.tex" | |
| 411 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 412 | session "HOL-ex" in ex = HOL + | 
| 48481 | 413 |   description {* Miscellaneous examples for Higher-Order Logic. *}
 | 
| 48679 | 414 | options [timeout = 3600, condition = ISABELLE_POLYML] | 
| 48481 | 415 | theories [document = false] | 
| 416 | "~~/src/HOL/Library/State_Monad" | |
| 417 | Code_Nat_examples | |
| 418 | "~~/src/HOL/Library/FuncSet" | |
| 419 | Eval_Examples | |
| 420 | Normalization_by_Evaluation | |
| 421 | Hebrew | |
| 422 | Chinese | |
| 423 | Serbian | |
| 424 | "~~/src/HOL/Library/FinFun_Syntax" | |
| 425 | theories | |
| 426 | Iff_Oracle | |
| 427 | Coercion_Examples | |
| 428 | Numeral_Representation | |
| 429 | Higher_Order_Logic | |
| 430 | Abstract_NAT | |
| 431 | Guess | |
| 432 | Binary | |
| 433 | Fundefs | |
| 434 | Induction_Schema | |
| 435 | LocaleTest2 | |
| 436 | Records | |
| 437 | While_Combinator_Example | |
| 438 | MonoidGroup | |
| 439 | BinEx | |
| 440 | Hex_Bin_Examples | |
| 441 | Antiquote | |
| 442 | Multiquote | |
| 443 | PER | |
| 444 | NatSum | |
| 445 | ThreeDivides | |
| 446 | Intuitionistic | |
| 447 | CTL | |
| 448 | Arith_Examples | |
| 449 | BT | |
| 450 | Tree23 | |
| 451 | MergeSort | |
| 452 | Lagrange | |
| 453 | Groebner_Examples | |
| 454 | MT | |
| 455 | Unification | |
| 456 | Primrec | |
| 457 | Tarski | |
| 458 | Classical | |
| 459 | Set_Theory | |
| 460 | Meson_Test | |
| 461 | Termination | |
| 462 | Coherent | |
| 463 | PresburgerEx | |
| 464 | ReflectionEx | |
| 465 | Sqrt | |
| 466 | Sqrt_Script | |
| 467 | Transfer_Ex | |
| 468 | Transfer_Int_Nat | |
| 469 | HarmonicSeries | |
| 470 | Refute_Examples | |
| 471 | Landau | |
| 472 | Execute_Choice | |
| 473 | Summation | |
| 474 | Gauge_Integration | |
| 475 | Dedekind_Real | |
| 476 | Quicksort | |
| 477 | Birthday_Paradox | |
| 478 | List_to_Set_Comprehension_Examples | |
| 479 | Seq | |
| 480 | Simproc_Tests | |
| 481 | Executable_Relation | |
| 482 | FinFunPred | |
| 483 | Set_Comprehension_Pointfree_Tests | |
| 484 | Parallel_Example | |
| 485 | theories SVC_Oracle | |
| 48690 | 486 | theories [condition = SVC_HOME] | 
| 487 | svc_test | |
| 48481 | 488 | theories [condition = ZCHAFF_HOME] | 
| 489 | (*requires zChaff (or some other reasonably fast SAT solver)*) | |
| 490 | Sudoku | |
| 491 | (* FIXME | |
| 492 | (*requires a proof-generating SAT solver (zChaff or MiniSAT)*) | |
| 493 | (*global side-effects ahead!*) | |
| 494 | try use_thy "SAT_Examples"; (* FIXME try!? (not really a proper test) *) | |
| 495 | *) | |
| 496 | files "document/root.bib" "document/root.tex" | |
| 497 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 498 | session "HOL-Isar_Examples" in Isar_Examples = HOL + | 
| 48481 | 499 |   description {* Miscellaneous Isabelle/Isar examples for Higher-Order Logic. *}
 | 
| 500 | theories [document = false] | |
| 501 | "~~/src/HOL/Library/Lattice_Syntax" | |
| 502 | "../Number_Theory/Primes" | |
| 503 | theories | |
| 504 | Basic_Logic | |
| 505 | Cantor | |
| 506 | Drinker | |
| 507 | Expr_Compiler | |
| 508 | Fibonacci | |
| 509 | Group | |
| 510 | Group_Context | |
| 511 | Group_Notepad | |
| 512 | Hoare_Ex | |
| 513 | Knaster_Tarski | |
| 514 | Mutilated_Checkerboard | |
| 515 | Nested_Datatype | |
| 516 | Peirce | |
| 517 | Puzzle | |
| 518 | Summation | |
| 519 | files | |
| 520 | "document/root.bib" | |
| 521 | "document/root.tex" | |
| 522 | "document/style.tex" | |
| 523 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 524 | session "HOL-SET_Protocol" in SET_Protocol = HOL + | 
| 48481 | 525 | options [document_graph] | 
| 526 | theories [document = false] "~~/src/HOL/Library/Nat_Bijection" | |
| 527 | theories SET_Protocol | |
| 528 | files "document/root.tex" | |
| 529 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 530 | session "HOL-Matrix_LP" in Matrix_LP = HOL + | 
| 48481 | 531 | options [document_graph] | 
| 532 | theories Cplex | |
| 533 | files "document/root.tex" | |
| 534 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 535 | session "HOL-TLA" in TLA = HOL + | 
| 48481 | 536 |   description {* The Temporal Logic of Actions *}
 | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 537 | options [document = false] | 
| 48481 | 538 | theories TLA | 
| 539 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 540 | session "HOL-TLA-Inc" in "TLA/Inc" = "HOL-TLA" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 541 | options [document = false] | 
| 48481 | 542 | theories Inc | 
| 543 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 544 | session "HOL-TLA-Buffer" in "TLA/Buffer" = "HOL-TLA" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 545 | options [document = false] | 
| 48481 | 546 | theories DBuffer | 
| 547 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 548 | session "HOL-TLA-Memory" in "TLA/Memory" = "HOL-TLA" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 549 | options [document = false] | 
| 48481 | 550 | theories MemoryImplementation | 
| 551 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 552 | session "HOL-TPTP" in TPTP = HOL + | 
| 48481 | 553 |   description {*
 | 
| 554 | Author: Jasmin Blanchette, TU Muenchen | |
| 555 | Author: Nik Sultana, University of Cambridge | |
| 556 | Copyright 2011 | |
| 557 | ||
| 558 | TPTP-related extensions. | |
| 559 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 560 | options [document = false] | 
| 48481 | 561 | theories | 
| 562 | ATP_Theory_Export | |
| 563 | MaSh_Eval | |
| 564 | MaSh_Export | |
| 565 | TPTP_Interpret | |
| 566 | THF_Arith | |
| 567 | theories [proofs = 0] (* FIXME !? *) | |
| 568 | ATP_Problem_Import | |
| 569 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 570 | session "HOL-Multivariate_Analysis" in Multivariate_Analysis = HOL + | 
| 48481 | 571 | options [document_graph] | 
| 572 | theories | |
| 573 | Multivariate_Analysis | |
| 574 | Determinants | |
| 575 | files | |
| 576 | "Integration.certs" | |
| 577 | "document/root.tex" | |
| 578 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 579 | session "HOL-Probability" in "Probability" = "HOL-Multivariate_Analysis" + | 
| 48617 | 580 | options [document_graph] | 
| 48481 | 581 | theories [document = false] | 
| 582 | "~~/src/HOL/Library/Countable" | |
| 583 | "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits" | |
| 584 | "~~/src/HOL/Library/Permutation" | |
| 585 | theories | |
| 586 | Probability | |
| 587 | "ex/Dining_Cryptographers" | |
| 588 | "ex/Koepf_Duermuth_Countermeasure" | |
| 589 | files "document/root.tex" | |
| 590 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 591 | session "HOL-Nominal" (main) in Nominal = HOL + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 592 | options [document = false] | 
| 48481 | 593 | theories Nominal | 
| 594 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 595 | session "HOL-Nominal-Examples" in "Nominal/Examples" = "HOL-Nominal" + | 
| 48679 | 596 | options [timeout = 3600, condition = ISABELLE_POLYML, document = false] | 
| 48481 | 597 | theories Nominal_Examples | 
| 598 | theories [quick_and_dirty] VC_Condition | |
| 599 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 600 | session "HOL-Ordinals_and_Cardinals-Base" in Ordinals_and_Cardinals = HOL + | 
| 48978 | 601 |   description {* Ordinals and Cardinals, Base Theories *}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 602 | options [document = false] | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 603 | theories Cardinal_Arithmetic | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 604 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 605 | session "HOL-Ordinals_and_Cardinals" in Ordinals_and_Cardinals = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 606 | "HOL-Ordinals_and_Cardinals-Base" + | 
| 48978 | 607 |   description {* Ordinals and Cardinals, Full Theories *}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 608 | theories Cardinal_Order_Relation | 
| 48984 
f51d4a302962
do not hardwire document output options -- to be provided by the user;
 wenzelm parents: 
48978diff
changeset | 609 | files | 
| 
f51d4a302962
do not hardwire document output options -- to be provided by the user;
 wenzelm parents: 
48978diff
changeset | 610 | "document/intro.tex" | 
| 
f51d4a302962
do not hardwire document output options -- to be provided by the user;
 wenzelm parents: 
48978diff
changeset | 611 | "document/root.tex" | 
| 
f51d4a302962
do not hardwire document output options -- to be provided by the user;
 wenzelm parents: 
48978diff
changeset | 612 | "document/root.bib" | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 613 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 614 | session "HOL-Codatatype" in Codatatype = "HOL-Ordinals_and_Cardinals-Base" + | 
| 48978 | 615 |   description {* New (Co)datatype Package *}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 616 | options [document = false] | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 617 | theories Codatatype | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 618 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 619 | session "HOL-Codatatype-Examples" in "Codatatype/Examples" = "HOL-Codatatype" + | 
| 48978 | 620 |   description {* Examples for the New (Co)datatype Package *}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 621 | options [document = false] | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 622 | theories | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 623 | HFset | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 624 | Lambda_Term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 625 | Process | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 626 | TreeFsetI | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 627 | "Infinite_Derivation_Trees/Gram_Lang" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 628 | "Infinite_Derivation_Trees/Parallel" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 629 | Stream | 
| 49110 
2e43fb45b91b
eliminated obsolete "parallel_proofs = 0" restriction (cf. 0e5b859e1c91)
 traytel parents: 
49077diff
changeset | 630 | theories [condition = ISABELLE_FULL_TEST] | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 631 | Misc_Codata | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 632 | Misc_Data | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: 
48932diff
changeset | 633 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 634 | session "HOL-Word" in Word = HOL + | 
| 48481 | 635 | options [document_graph] | 
| 636 | theories Word | |
| 637 | files "document/root.bib" "document/root.tex" | |
| 638 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 639 | session "HOL-Word-Examples" in "Word/Examples" = "HOL-Word" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 640 | options [document = false] | 
| 48481 | 641 | theories WordExamples | 
| 642 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 643 | session "HOL-Statespace" in Statespace = HOL + | 
| 48481 | 644 | theories StateSpaceEx | 
| 645 | files "document/root.tex" | |
| 646 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 647 | session "HOL-NSA" in NSA = HOL + | 
| 48481 | 648 | options [document_graph] | 
| 649 | theories Hypercomplex | |
| 650 | files "document/root.tex" | |
| 651 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 652 | session "HOL-NSA-Examples" in "NSA/Examples" = "HOL-NSA" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 653 | options [document = false] | 
| 48481 | 654 | theories NSPrimes | 
| 655 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 656 | session "HOL-Mirabelle" in Mirabelle = HOL + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 657 | options [document = false] | 
| 48481 | 658 | theories Mirabelle_Test | 
| 48589 
fb446a780d50
separate session HOL-Mirabelle-ex -- cannot run isolated shell scripts within build tool;
 wenzelm parents: 
48588diff
changeset | 659 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 660 | session "HOL-Mirabelle-ex" in "Mirabelle/ex" = "HOL-Mirabelle" + | 
| 48589 
fb446a780d50
separate session HOL-Mirabelle-ex -- cannot run isolated shell scripts within build tool;
 wenzelm parents: 
48588diff
changeset | 661 | theories Ex | 
| 48481 | 662 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 663 | session "HOL-Word-SMT_Examples" in SMT_Examples = "HOL-Word" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 664 | options [document = false, quick_and_dirty] | 
| 48481 | 665 | theories | 
| 666 | SMT_Tests | |
| 667 | SMT_Examples | |
| 668 | SMT_Word_Examples | |
| 669 | files | |
| 670 | "SMT_Examples.certs" | |
| 671 | "SMT_Tests.certs" | |
| 672 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 673 | session "HOL-Boogie" in "Boogie" = "HOL-Word" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 674 | options [document = false] | 
| 48481 | 675 | theories Boogie | 
| 676 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 677 | session "HOL-Boogie-Examples" in "Boogie/Examples" = "HOL-Boogie" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 678 | options [document = false] | 
| 48481 | 679 | theories | 
| 680 | Boogie_Max_Stepwise | |
| 681 | Boogie_Max | |
| 682 | Boogie_Dijkstra | |
| 683 | VCC_Max | |
| 684 | files | |
| 48493 | 685 | "Boogie_Dijkstra.b2i" | 
| 48481 | 686 | "Boogie_Dijkstra.certs" | 
| 48493 | 687 | "Boogie_Max.b2i" | 
| 48481 | 688 | "Boogie_Max.certs" | 
| 48493 | 689 | "VCC_Max.b2i" | 
| 48481 | 690 | "VCC_Max.certs" | 
| 691 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 692 | session "HOL-SPARK" in "SPARK" = "HOL-Word" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 693 | options [document = false] | 
| 48481 | 694 | theories SPARK | 
| 695 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 696 | session "HOL-SPARK-Examples" in "SPARK/Examples" = "HOL-SPARK" + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 697 | options [document = false] | 
| 48481 | 698 | theories | 
| 699 | "Gcd/Greatest_Common_Divisor" | |
| 700 | ||
| 701 | "Liseq/Longest_Increasing_Subsequence" | |
| 702 | ||
| 703 | "RIPEMD-160/F" | |
| 704 | "RIPEMD-160/Hash" | |
| 705 | "RIPEMD-160/K_L" | |
| 706 | "RIPEMD-160/K_R" | |
| 707 | "RIPEMD-160/R_L" | |
| 708 | "RIPEMD-160/Round" | |
| 709 | "RIPEMD-160/R_R" | |
| 710 | "RIPEMD-160/S_L" | |
| 711 | "RIPEMD-160/S_R" | |
| 712 | ||
| 713 | "Sqrt/Sqrt" | |
| 714 | files | |
| 715 | "Gcd/greatest_common_divisor/g_c_d.fdl" | |
| 716 | "Gcd/greatest_common_divisor/g_c_d.rls" | |
| 717 | "Gcd/greatest_common_divisor/g_c_d.siv" | |
| 718 | "Liseq/liseq/liseq_length.fdl" | |
| 719 | "Liseq/liseq/liseq_length.rls" | |
| 720 | "Liseq/liseq/liseq_length.siv" | |
| 721 | "RIPEMD-160/rmd/f.fdl" | |
| 722 | "RIPEMD-160/rmd/f.rls" | |
| 723 | "RIPEMD-160/rmd/f.siv" | |
| 724 | "RIPEMD-160/rmd/hash.fdl" | |
| 725 | "RIPEMD-160/rmd/hash.rls" | |
| 726 | "RIPEMD-160/rmd/hash.siv" | |
| 727 | "RIPEMD-160/rmd/k_l.fdl" | |
| 728 | "RIPEMD-160/rmd/k_l.rls" | |
| 729 | "RIPEMD-160/rmd/k_l.siv" | |
| 730 | "RIPEMD-160/rmd/k_r.fdl" | |
| 731 | "RIPEMD-160/rmd/k_r.rls" | |
| 732 | "RIPEMD-160/rmd/k_r.siv" | |
| 733 | "RIPEMD-160/rmd/r_l.fdl" | |
| 734 | "RIPEMD-160/rmd/r_l.rls" | |
| 735 | "RIPEMD-160/rmd/r_l.siv" | |
| 736 | "RIPEMD-160/rmd/round.fdl" | |
| 737 | "RIPEMD-160/rmd/round.rls" | |
| 738 | "RIPEMD-160/rmd/round.siv" | |
| 739 | "RIPEMD-160/rmd/r_r.fdl" | |
| 740 | "RIPEMD-160/rmd/r_r.rls" | |
| 741 | "RIPEMD-160/rmd/r_r.siv" | |
| 742 | "RIPEMD-160/rmd/s_l.fdl" | |
| 743 | "RIPEMD-160/rmd/s_l.rls" | |
| 744 | "RIPEMD-160/rmd/s_l.siv" | |
| 745 | "RIPEMD-160/rmd/s_r.fdl" | |
| 746 | "RIPEMD-160/rmd/s_r.rls" | |
| 747 | "RIPEMD-160/rmd/s_r.siv" | |
| 748 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 749 | session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" + | 
| 48486 | 750 | options [show_question_marks = false] | 
| 48481 | 751 | theories | 
| 752 | Example_Verification | |
| 753 | VC_Principles | |
| 754 | Reference | |
| 755 | Complex_Types | |
| 756 | files | |
| 757 | "complex_types_app/initialize.fdl" | |
| 758 | "complex_types_app/initialize.rls" | |
| 759 | "complex_types_app/initialize.siv" | |
| 760 | "document/complex_types.ads" | |
| 761 | "document/complex_types_app.adb" | |
| 762 | "document/complex_types_app.ads" | |
| 763 | "document/Gcd.adb" | |
| 764 | "document/Gcd.ads" | |
| 765 | "document/intro.tex" | |
| 766 | "document/loop_invariant.adb" | |
| 767 | "document/loop_invariant.ads" | |
| 768 | "document/root.bib" | |
| 769 | "document/root.tex" | |
| 770 | "document/Simple_Gcd.adb" | |
| 771 | "document/Simple_Gcd.ads" | |
| 772 | "loop_invariant/proc1.fdl" | |
| 773 | "loop_invariant/proc1.rls" | |
| 774 | "loop_invariant/proc1.siv" | |
| 775 | "loop_invariant/proc2.fdl" | |
| 776 | "loop_invariant/proc2.rls" | |
| 777 | "loop_invariant/proc2.siv" | |
| 778 | "simple_greatest_common_divisor/g_c_d.fdl" | |
| 779 | "simple_greatest_common_divisor/g_c_d.rls" | |
| 780 | "simple_greatest_common_divisor/g_c_d.siv" | |
| 781 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 782 | session "HOL-Mutabelle" in Mutabelle = HOL + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 783 | options [document = false] | 
| 48481 | 784 | theories MutabelleExtra | 
| 785 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 786 | session "HOL-Quickcheck_Examples" in Quickcheck_Examples = HOL + | 
| 48679 | 787 | options [timeout = 3600, document = false] | 
| 48588 | 788 | theories | 
| 48690 | 789 | Quickcheck_Examples | 
| 790 | (* FIXME | |
| 791 | Quickcheck_Lattice_Examples | |
| 792 | Completeness | |
| 793 | Quickcheck_Interfaces | |
| 794 | Hotel_Example *) | |
| 48598 | 795 | theories [condition = ISABELLE_GHC] | 
| 796 | Quickcheck_Narrowing_Examples | |
| 48588 | 797 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 798 | session "HOL-Quickcheck_Benchmark" in Quickcheck_Benchmark = HOL + | 
| 48635 | 799 | theories [condition = ISABELLE_FULL_TEST, quick_and_dirty] | 
| 48618 
1f7e068b4613
moved another larger quickcheck example to Quickcheck_Benchmark
 bulwahn parents: 
48614diff
changeset | 800 | Find_Unused_Assms_Examples | 
| 
1f7e068b4613
moved another larger quickcheck example to Quickcheck_Benchmark
 bulwahn parents: 
48614diff
changeset | 801 | Needham_Schroeder_No_Attacker_Example | 
| 
1f7e068b4613
moved another larger quickcheck example to Quickcheck_Benchmark
 bulwahn parents: 
48614diff
changeset | 802 | Needham_Schroeder_Guided_Attacker_Example | 
| 48690 | 803 | Needham_Schroeder_Unguided_Attacker_Example | 
| 48481 | 804 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 805 | session "HOL-Quotient_Examples" in Quotient_Examples = HOL + | 
| 48481 | 806 |   description {*
 | 
| 807 | Author: Cezary Kaliszyk and Christian Urban | |
| 808 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 809 | options [document = false] | 
| 48481 | 810 | theories | 
| 811 | DList | |
| 812 | FSet | |
| 813 | Quotient_Int | |
| 814 | Quotient_Message | |
| 815 | Lift_FSet | |
| 816 | Lift_Set | |
| 817 | Lift_Fun | |
| 818 | Quotient_Rat | |
| 819 | Lift_DList | |
| 820 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 821 | session "HOL-Predicate_Compile_Examples" in Predicate_Compile_Examples = HOL + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 822 | options [document = false] | 
| 48690 | 823 | theories | 
| 48481 | 824 | Examples | 
| 825 | Predicate_Compile_Tests | |
| 48690 | 826 | (* FIXME | 
| 827 | Predicate_Compile_Quickcheck_Examples -- should be added again soon (since 21-Oct-2010) *) | |
| 48481 | 828 | Specialisation_Examples | 
| 48690 | 829 | (* FIXME since 21-Jul-2011 | 
| 830 | Hotel_Example_Small_Generator | |
| 831 | IMP_1 | |
| 832 | IMP_2 | |
| 833 | IMP_3 | |
| 834 | IMP_4 *) | |
| 835 | theories [condition = "ISABELLE_SWIPL"] (* FIXME: *or* ISABELLE_YAP (??) *) | |
| 836 | Code_Prolog_Examples | |
| 837 | Context_Free_Grammar_Example | |
| 838 | Hotel_Example_Prolog | |
| 839 | Lambda_Example | |
| 840 | List_Examples | |
| 841 | theories [condition = "ISABELLE_SWIPL", quick_and_dirty] (* FIXME: *or* ISABELLE_YAP (??) *) | |
| 842 | Reg_Exp_Example | |
| 48481 | 843 | |
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 844 | session HOLCF (main) in HOLCF = HOL + | 
| 48338 | 845 |   description {*
 | 
| 846 | Author: Franz Regensburger | |
| 847 | Author: Brian Huffman | |
| 848 | ||
| 849 | HOLCF -- a semantic extension of HOL by the LCF logic. | |
| 850 | *} | |
| 851 | options [document_graph] | |
| 48470 
7483aa690b4f
clarified "document" again, eliminated redundant "no_document";
 wenzelm parents: 
48458diff
changeset | 852 | theories [document = false] | 
| 48338 | 853 | "~~/src/HOL/Library/Nat_Bijection" | 
| 854 | "~~/src/HOL/Library/Countable" | |
| 48481 | 855 | theories | 
| 856 | Plain_HOLCF | |
| 857 | Fixrec | |
| 858 | HOLCF | |
| 859 | files "document/root.tex" | |
| 860 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 861 | session "HOLCF-Tutorial" in "HOLCF/Tutorial" = HOLCF + | 
| 48481 | 862 | theories | 
| 863 | Domain_ex | |
| 864 | Fixrec_ex | |
| 865 | New_Domain | |
| 866 | files "document/root.tex" | |
| 867 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 868 | session "HOLCF-Library" in "HOLCF/Library" = HOLCF + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 869 | options [document = false] | 
| 48481 | 870 | theories HOLCF_Library | 
| 871 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 872 | session "HOLCF-IMP" in "HOLCF/IMP" = HOLCF + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 873 | options [document = false] | 
| 48481 | 874 | theories HoareEx | 
| 48338 | 875 | files "document/root.tex" | 
| 876 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 877 | session "HOLCF-ex" in "HOLCF/ex" = HOLCF + | 
| 48481 | 878 |   description {* Misc HOLCF examples *}
 | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 879 | options [document = false] | 
| 48481 | 880 | theories | 
| 881 | Dnat | |
| 882 | Dagstuhl | |
| 883 | Focus_ex | |
| 884 | Fix2 | |
| 885 | Hoare | |
| 886 | Concurrency_Monad | |
| 887 | Loop | |
| 888 | Powerdomain_ex | |
| 889 | Domain_Proofs | |
| 890 | Letrec | |
| 891 | Pattern_Match | |
| 892 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 893 | session "HOLCF-FOCUS" in "HOLCF/FOCUS" = HOLCF + | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 894 | options [document = false] | 
| 48481 | 895 | theories | 
| 896 | Fstreams | |
| 897 | FOCUS | |
| 898 | Buffer_adm | |
| 899 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 900 | session IOA in "HOLCF/IOA" = HOLCF + | 
| 48481 | 901 |   description {*
 | 
| 902 | Author: Olaf Mueller | |
| 903 | ||
| 904 | Formalization of a semantic model of I/O-Automata. | |
| 905 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 906 | options [document = false] | 
| 48481 | 907 | theories "meta_theory/Abstraction" | 
| 908 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 909 | session "IOA-ABP" in "HOLCF/IOA/ABP" = IOA + | 
| 48481 | 910 |   description {*
 | 
| 911 | Author: Olaf Mueller | |
| 912 | ||
| 913 | The Alternating Bit Protocol performed in I/O-Automata. | |
| 914 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 915 | options [document = false] | 
| 48481 | 916 | theories Correctness | 
| 917 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 918 | session "IOA-NTP" in "HOLCF/IOA/NTP" = IOA + | 
| 48481 | 919 |   description {*
 | 
| 920 | Author: Tobias Nipkow & Konrad Slind | |
| 921 | ||
| 922 | A network transmission protocol, performed in the | |
| 923 | I/O automata formalization by Olaf Mueller. | |
| 924 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 925 | options [document = false] | 
| 48481 | 926 | theories Correctness | 
| 927 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 928 | session "IOA-Storage" in "HOLCF/IOA/Storage" = IOA + | 
| 48481 | 929 |   description {*
 | 
| 930 | Author: Olaf Mueller | |
| 931 | ||
| 932 | Memory storage case study. | |
| 933 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 934 | options [document = false] | 
| 48481 | 935 | theories Correctness | 
| 936 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 937 | session "IOA-ex" in "HOLCF/IOA/ex" = IOA + | 
| 48481 | 938 |   description {*
 | 
| 939 | Author: Olaf Mueller | |
| 940 | *} | |
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 941 | options [document = false] | 
| 48481 | 942 | theories | 
| 943 | TrivEx | |
| 944 | TrivEx2 | |
| 945 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 946 | session "HOL-Datatype_Benchmark" in Datatype_Benchmark = HOL + | 
| 48481 | 947 |   description {* Some rather large datatype examples (from John Harrison). *}
 | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 948 | options [document = false] | 
| 48635 | 949 | theories [condition = ISABELLE_FULL_TEST, timing] | 
| 48481 | 950 | Brackin | 
| 951 | Instructions | |
| 952 | SML | |
| 953 | Verilog | |
| 954 | ||
| 48738 
f8c1a5b9488f
simplified session specifications: names are taken verbatim and current directory is default;
 wenzelm parents: 
48721diff
changeset | 955 | session "HOL-Record_Benchmark" in Record_Benchmark = HOL + | 
| 48481 | 956 |   description {* Some benchmark on large record. *}
 | 
| 48483 
9bfb6978eb80
more explicit document = false to reduce warnings;
 wenzelm parents: 
48481diff
changeset | 957 | options [document = false] | 
| 48635 | 958 | theories [condition = ISABELLE_FULL_TEST, timing] | 
| 48481 | 959 | Record_Benchmark | 
| 960 |