| author | wenzelm |
| Sun, 29 Sep 2024 19:51:23 +0200 | |
| changeset 81004 | 07ad0b407d38 |
| parent 80914 | d97fdabd9e2b |
| child 81136 | 2b949a3bfaac |
| 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:
61670
diff
changeset
|
4 |
"HOL-Library.Dlist" |
|
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
61670
diff
changeset
|
5 |
"HOL-Library.RBT" |
|
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
61670
diff
changeset
|
6 |
"HOL-Library.Mapping" |
|
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
61670
diff
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) |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69605
diff
changeset
|
14 |
"_constrain" :: "logic => type => logic" (\<open>_::_\<close> [4, 0] 3) |
|
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69605
diff
changeset
|
15 |
"_constrain" :: "prop' => type => prop'" (\<open>_::_\<close> [4, 0] 3) |
| 61143 | 16 |
|
17 |
syntax (output) |
|
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69605
diff
changeset
|
18 |
"_constrain" :: "logic => type => logic" (\<open>_ :: _\<close> [4, 0] 3) |
|
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69605
diff
changeset
|
19 |
"_constrain" :: "prop' => type => prop'" (\<open>_ :: _\<close> [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:
42358
diff
changeset
|
23 |
declare [[names_unique = false]] |
| 28213 | 24 |
|
25 |
end |