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