| 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
 | 
|  |      9 | imports Main
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | (* LOGIC *)
 | 
| 19674 |     13 | const_syntax (latex output)
 | 
|  |     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 *)
 | 
| 19674 |     55 | const_syntax (latex)
 | 
|  |     56 |   Cons  ("_\<cdot>/_" [66,65] 65)
 | 
| 15469 |     57 | 
 | 
|  |     58 | (* length *)
 | 
| 19674 |     59 | const_syntax (latex output)
 | 
|  |     60 |   length  ("|_|")
 | 
| 15469 |     61 | 
 | 
|  |     62 | (* nth *)
 | 
| 19674 |     63 | const_syntax (latex output)
 | 
|  |     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 | 
 | 
|  |     82 | syntax (IfThen output)
 | 
|  |     83 |   "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
 | 
|  |     84 |   ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
 | 
|  |     85 |   "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
 | 
|  |     86 |   ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
 | 
|  |     87 |   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\rmfamily\upshape\normalsize \,>and\<^raw:\,}>/ _")
 | 
|  |     88 |   "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
 | 
|  |     89 | 
 | 
|  |     90 | syntax (IfThenNoBox output)
 | 
|  |     91 |   "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
 | 
|  |     92 |   ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
 | 
|  |     93 |   "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
 | 
|  |     94 |   ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
 | 
|  |     95 |   "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\rmfamily\upshape\normalsize \,>and\<^raw:\,}>/ _")
 | 
|  |     96 |   "_asm" :: "prop \<Rightarrow> asms" ("_")
 | 
|  |     97 | 
 | 
|  |     98 | end
 | 
|  |     99 | (*>*) |