author | haftmann |
Mon, 10 Dec 2007 11:24:12 +0100 | |
changeset 25595 | 6c48275f9c76 |
parent 25467 | bba589a88022 |
child 27368 | 9f90ac19e32b |
permissions | -rw-r--r-- |
15469 | 1 |
(* Title: HOL/Library/LaTeXsugar.thy |
2 |
ID: $Id$ |
|
3 |
Author: Gerwin Klain, Tobias Nipkow, Norbert Schirmer |
|
4 |
Copyright 2005 NICTA and TUM |
|
5 |
*) |
|
6 |
||
7 |
(*<*) |
|
8 |
theory LaTeXsugar |
|
25595 | 9 |
imports List |
15469 | 10 |
begin |
11 |
||
12 |
(* LOGIC *) |
|
21210 | 13 |
notation (latex output) |
19674 | 14 |
If ("(\<^raw:\textsf{>if\<^raw:}> (_)/ \<^raw:\textsf{>then\<^raw:}> (_)/ \<^raw:\textsf{>else\<^raw:}> (_))" 10) |
15 |
||
15469 | 16 |
syntax (latex output) |
17 |
||
18 |
"_Let" :: "[letbinds, 'a] => 'a" |
|
19 |
("(\<^raw:\textsf{>let\<^raw:}> (_)/ \<^raw:\textsf{>in\<^raw:}> (_))" 10) |
|
20 |
||
21 |
"_case_syntax":: "['a, cases_syn] => 'b" |
|
22 |
("(\<^raw:\textsf{>case\<^raw:}> _ \<^raw:\textsf{>of\<^raw:}>/ _)" 10) |
|
23 |
||
16110 | 24 |
(* should become standard syntax once x-symbols supports it *) |
25 |
syntax (latex) |
|
26 |
nexists :: "('a => bool) => bool" (binder "\<nexists>" 10) |
|
27 |
translations |
|
28 |
"\<nexists>x. P" <= "\<not>(\<exists>x. P)" |
|
29 |
||
15469 | 30 |
(* SETS *) |
31 |
||
32 |
(* empty set *) |
|
33 |
syntax (latex output) |
|
34 |
"_emptyset" :: "'a set" ("\<emptyset>") |
|
35 |
translations |
|
36 |
"_emptyset" <= "{}" |
|
37 |
||
38 |
(* insert *) |
|
39 |
translations |
|
40 |
"{x} \<union> A" <= "insert x A" |
|
15690 | 41 |
"{x,y}" <= "{x} \<union> {y}" |
15469 | 42 |
"{x,y} \<union> A" <= "{x} \<union> ({y} \<union> A)" |
15476 | 43 |
"{x}" <= "{x} \<union> _emptyset" |
15469 | 44 |
|
45 |
(* set comprehension *) |
|
46 |
syntax (latex output) |
|
47 |
"_Collect" :: "pttrn => bool => 'a set" ("(1{_ | _})") |
|
48 |
translations |
|
49 |
"_Collect p P" <= "{p. P}" |
|
15690 | 50 |
"_Collect p P" <= "{p|xs. P}" |
15469 | 51 |
|
52 |
(* LISTS *) |
|
53 |
||
54 |
(* Cons *) |
|
21210 | 55 |
notation (latex) |
19674 | 56 |
Cons ("_\<cdot>/_" [66,65] 65) |
15469 | 57 |
|
58 |
(* length *) |
|
21210 | 59 |
notation (latex output) |
19674 | 60 |
length ("|_|") |
15469 | 61 |
|
62 |
(* nth *) |
|
21210 | 63 |
notation (latex output) |
19674 | 64 |
nth ("_\<^raw:\ensuremath{_{[\mathit{>_\<^raw:}]}}>" [1000,0] 1000) |
15469 | 65 |
|
66 |
(* DUMMY *) |
|
67 |
consts DUMMY :: 'a ("\<^raw:\_>") |
|
68 |
||
69 |
(* THEOREMS *) |
|
70 |
syntax (Rule output) |
|
71 |
"==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop" |
|
72 |
("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>") |
|
73 |
||
74 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
|
75 |
("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>") |
|
76 |
||
77 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" |
|
78 |
("\<^raw:\mbox{>_\<^raw:}\\>/ _") |
|
79 |
||
80 |
"_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>") |
|
81 |
||
22328
cc403d881873
added print-mode Axiom to print theorems without premises with a rule on top.
schirmer
parents:
21210
diff
changeset
|
82 |
syntax (Axiom output) |
cc403d881873
added print-mode Axiom to print theorems without premises with a rule on top.
schirmer
parents:
21210
diff
changeset
|
83 |
"Trueprop" :: "bool \<Rightarrow> prop" |
cc403d881873
added print-mode Axiom to print theorems without premises with a rule on top.
schirmer
parents:
21210
diff
changeset
|
84 |
("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>") |
cc403d881873
added print-mode Axiom to print theorems without premises with a rule on top.
schirmer
parents:
21210
diff
changeset
|
85 |
|
15469 | 86 |
syntax (IfThen output) |
87 |
"==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop" |
|
25467 | 88 |
("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.") |
15469 | 89 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
25467 | 90 |
("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.") |
91 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _") |
|
15469 | 92 |
"_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>") |
93 |
||
94 |
syntax (IfThenNoBox output) |
|
95 |
"==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop" |
|
25467 | 96 |
("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.") |
15469 | 97 |
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop" |
25467 | 98 |
("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.") |
99 |
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _") |
|
15469 | 100 |
"_asm" :: "prop \<Rightarrow> asms" ("_") |
101 |
||
102 |
end |
|
103 |
(*>*) |