| author | wenzelm | 
| Thu, 27 Feb 2020 12:58:03 +0100 | |
| changeset 71489 | e8da4a8d364a | 
| parent 69605 | a96320074298 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 28213 | 1 | theory Setup | 
| 39066 | 2 | imports | 
| 3 | Complex_Main | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
61670diff
changeset | 4 | "HOL-Library.Dlist" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
61670diff
changeset | 5 | "HOL-Library.RBT" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
61670diff
changeset | 6 | "HOL-Library.Mapping" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
61670diff
changeset | 7 | "HOL-Library.IArray" | 
| 28213 | 8 | begin | 
| 9 | ||
| 69605 | 10 | ML_file \<open>../antiquote_setup.ML\<close> | 
| 11 | ML_file \<open>../more_antiquote.ML\<close> | |
| 48891 | 12 | |
| 61143 | 13 | no_syntax (output) | 
| 14 |   "_constrain" :: "logic => type => logic"  ("_::_" [4, 0] 3)
 | |
| 15 |   "_constrain" :: "prop' => type => prop'"  ("_::_" [4, 0] 3)
 | |
| 16 | ||
| 17 | syntax (output) | |
| 18 |   "_constrain" :: "logic => type => logic"  ("_ :: _" [4, 0] 3)
 | |
| 19 |   "_constrain" :: "prop' => type => prop'"  ("_ :: _" [4, 0] 3)
 | |
| 38503 | 20 | |
| 59324 | 21 | declare [[default_code_width = 74]] | 
| 34071 | 22 | |
| 42669 
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
 wenzelm parents: 
42358diff
changeset | 23 | declare [[names_unique = false]] | 
| 28213 | 24 | |
| 25 | end |