author | paulson <lp15@cam.ac.uk> |
Mon, 28 Aug 2017 20:33:08 +0100 | |
changeset 66537 | e2249cd6df67 |
parent 63935 | aa1fe1103ab8 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
15469 | 1 |
(* Title: HOL/Library/OptionalSugar.thy |
55052 | 2 |
Author: Gerwin Klein, Tobias Nipkow |
15469 | 3 |
Copyright 2005 NICTA and TUM |
4 |
*) |
|
15476 | 5 |
(*<*) |
15469 | 6 |
theory OptionalSugar |
30663
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
haftmann
parents:
30503
diff
changeset
|
7 |
imports Complex_Main LaTeXsugar |
15469 | 8 |
begin |
9 |
||
61645 | 10 |
(* displaying theorems with conjunctions between premises: *) |
11 |
translations |
|
12 |
("prop") "P \<and> Q \<Longrightarrow> R" <= ("prop") "P \<Longrightarrow> Q \<Longrightarrow> R" |
|
13 |
||
30474 | 14 |
(* hiding set *) |
22835 | 15 |
translations |
30404
d03dd6301678
more robust treatment of (authentic) consts within translations;
wenzelm
parents:
29494
diff
changeset
|
16 |
"xs" <= "CONST set xs" |
22835 | 17 |
|
30474 | 18 |
(* hiding numeric conversions - embeddings only *) |
19 |
translations |
|
20 |
"n" <= "CONST of_nat n" |
|
21 |
"n" <= "CONST int n" |
|
35115 | 22 |
"n" <= "CONST real n" |
30474 | 23 |
"n" <= "CONST real_of_nat n" |
24 |
"n" <= "CONST real_of_int n" |
|
30502 | 25 |
"n" <= "CONST of_real n" |
26 |
"n" <= "CONST complex_of_real n" |
|
30474 | 27 |
|
15469 | 28 |
(* append *) |
29 |
syntax (latex output) |
|
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
61955
diff
changeset
|
30 |
"_appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^latex>\<open>\\isacharat\<close>" 65) |
15469 | 31 |
translations |
35115 | 32 |
"_appendL xs ys" <= "xs @ ys" |
33 |
"_appendL (_appendL xs ys) zs" <= "_appendL xs (_appendL ys zs)" |
|
15469 | 34 |
|
35 |
||
32891 | 36 |
(* deprecated, use thm with style instead, will be removed *) |
15469 | 37 |
(* aligning equations *) |
21210 | 38 |
notation (tab output) |
63935
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
61955
diff
changeset
|
39 |
"HOL.eq" ("(_) \<^latex>\<open>}\\putisatab\\isa{\\ \<close>=\<^latex>\<open>}\\putisatab\\isa{\<close> (_)" [50,49] 50) and |
aa1fe1103ab8
raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents:
61955
diff
changeset
|
40 |
"Pure.eq" ("(_) \<^latex>\<open>}\\putisatab\\isa{\\ \<close>\<equiv>\<^latex>\<open>}\\putisatab\\isa{\<close> (_)") |
15469 | 41 |
|
42 |
(* Let *) |
|
43 |
translations |
|
35250 | 44 |
"_bind (p, CONST DUMMY) e" <= "_bind p (CONST fst e)" |
45 |
"_bind (CONST DUMMY, p) e" <= "_bind p (CONST snd e)" |
|
15469 | 46 |
|
47 |
"_tuple_args x (_tuple_args y z)" == |
|
48 |
"_tuple_args x (_tuple_arg (_tuple y z))" |
|
49 |
||
35250 | 50 |
"_bind (CONST Some p) e" <= "_bind p (CONST the e)" |
51 |
"_bind (p # CONST DUMMY) e" <= "_bind p (CONST hd e)" |
|
52 |
"_bind (CONST DUMMY # p) e" <= "_bind p (CONST tl e)" |
|
15469 | 53 |
|
29494 | 54 |
(* type constraints with spacing *) |
61143 | 55 |
no_syntax (output) |
56 |
"_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) |
|
57 |
"_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) |
|
55038
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
58 |
|
61143 | 59 |
syntax (output) |
60 |
"_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) |
|
61 |
"_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) |
|
55038
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
62 |
|
29494 | 63 |
|
64 |
(* sorts as intersections *) |
|
55038
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
65 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61645
diff
changeset
|
66 |
syntax (output) |
55038
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
67 |
"_topsort" :: "sort" ("\<top>" 1000) |
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
68 |
"_sort" :: "classes => sort" ("'(_')" 1000) |
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
69 |
"_classes" :: "id => classes => classes" ("_ \<inter> _" 1000) |
f2179be64805
prefer Isar commands over old-fashioned ML (see also a189c6274c7a);
wenzelm
parents:
42297
diff
changeset
|
70 |
"_classes" :: "longid => classes => classes" ("_ \<inter> _" 1000) |
15469 | 71 |
|
15476 | 72 |
end |
73 |
(*>*) |