src/HOL/Library/OptionalSugar.thy
author wenzelm
Thu, 23 Jul 2009 16:53:00 +0200
changeset 32146 4937d9836824
parent 30663 0b6aff7451b2
child 32891 d403b99287ff
permissions -rw-r--r--
paramify_vars: Term_Subst.map_atypsT_same recovered coding conventions of this module;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15469
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Library/OptionalSugar.thy
nipkow
parents:
diff changeset
     2
    Author:     Gerwin Klain, Tobias Nipkow
nipkow
parents:
diff changeset
     3
    Copyright   2005 NICTA and TUM
nipkow
parents:
diff changeset
     4
*)
15476
b8cb20cc0c0b fixed bugs
nipkow
parents: 15469
diff changeset
     5
(*<*)
15469
nipkow
parents:
diff changeset
     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
nipkow
parents:
diff changeset
     8
begin
nipkow
parents:
diff changeset
     9
30474
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    10
(* hiding set *)
22835
37d3a984d730 added "set" supression
nipkow
parents: 21404
diff changeset
    11
translations
30404
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    12
  "xs" <= "CONST set xs"
22835
37d3a984d730 added "set" supression
nipkow
parents: 21404
diff changeset
    13
30474
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    14
(* hiding numeric conversions - embeddings only *)
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    15
translations
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    16
  "n" <= "CONST of_nat n"
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    17
  "n" <= "CONST int n"
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    18
  "n" <= "real n"
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    19
  "n" <= "CONST real_of_nat n"
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    20
  "n" <= "CONST real_of_int n"
30502
b80d2621caee hiding numeric coercions in LaTeX
nipkow
parents: 30474
diff changeset
    21
  "n" <= "CONST of_real n"
b80d2621caee hiding numeric coercions in LaTeX
nipkow
parents: 30474
diff changeset
    22
  "n" <= "CONST complex_of_real n"
30474
52e92009aacb optional latex sugar
nipkow
parents: 30404
diff changeset
    23
15469
nipkow
parents:
diff changeset
    24
(* append *)
nipkow
parents:
diff changeset
    25
syntax (latex output)
nipkow
parents:
diff changeset
    26
  "appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^raw:\isacharat>" 65)
nipkow
parents:
diff changeset
    27
translations
nipkow
parents:
diff changeset
    28
  "appendL xs ys" <= "xs @ ys" 
nipkow
parents:
diff changeset
    29
  "appendL (appendL xs ys) zs" <= "appendL xs (appendL ys zs)"
nipkow
parents:
diff changeset
    30
nipkow
parents:
diff changeset
    31
30503
201887dcea0a added comment
nipkow
parents: 30502
diff changeset
    32
(* deprecated, use thm_style instead, will be removed *)
15469
nipkow
parents:
diff changeset
    33
(* aligning equations *)
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 19674
diff changeset
    34
notation (tab output)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    35
  "op ="  ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
19674
22b635240905 const_syntax;
wenzelm
parents: 15476
diff changeset
    36
  "=="  ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")
15469
nipkow
parents:
diff changeset
    37
nipkow
parents:
diff changeset
    38
(* Let *)
nipkow
parents:
diff changeset
    39
translations 
30404
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    40
  "_bind (p,DUMMY) e" <= "_bind p (CONST fst e)"
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    41
  "_bind (DUMMY,p) e" <= "_bind p (CONST snd e)"
15469
nipkow
parents:
diff changeset
    42
nipkow
parents:
diff changeset
    43
  "_tuple_args x (_tuple_args y z)" ==
nipkow
parents:
diff changeset
    44
    "_tuple_args x (_tuple_arg (_tuple y z))"
nipkow
parents:
diff changeset
    45
30404
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    46
  "_bind (Some p) e" <= "_bind p (CONST the e)"
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    47
  "_bind (p#DUMMY) e" <= "_bind p (CONST hd e)"
d03dd6301678 more robust treatment of (authentic) consts within translations;
wenzelm
parents: 29494
diff changeset
    48
  "_bind (DUMMY#p) e" <= "_bind p (CONST tl e)"
15469
nipkow
parents:
diff changeset
    49
29494
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    50
(* type constraints with spacing *)
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    51
setup {*
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    52
let
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    53
  val typ = SimpleSyntax.read_typ;
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    54
  val typeT = Syntax.typeT;
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    55
  val spropT = Syntax.spropT;
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    56
in
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    57
  Sign.del_modesyntax_i (Symbol.xsymbolsN, false) [
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    58
    ("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    59
    ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_\<Colon>_", [4, 0], 3))]
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    60
  #> Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    61
      ("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    62
      ("_constrain", [spropT, typeT] ---> spropT, Mixfix ("_ \<Colon> _", [4, 0], 3))]
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    63
end
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    64
*}
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    65
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    66
(* sorts as intersections *)
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    67
setup {*
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    68
let
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    69
  val sortT = Type ("sort", []); (*FIXME*)
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    70
  val classesT = Type ("classes", []); (*FIXME*)
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    71
in
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    72
  Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    73
    ("_topsort", sortT, Mixfix ("\<top>", [], Syntax.max_pri)),
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    74
    ("_sort", classesT --> sortT, Mixfix ("'(_')", [], Syntax.max_pri)),
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    75
    ("_classes", Lexicon.idT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], Syntax.max_pri)),
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    76
    ("_classes", Lexicon.longidT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], Syntax.max_pri))
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    77
  ]
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    78
end
a189c6274c7a type constraints and sort intersection
haftmann
parents: 22835
diff changeset
    79
*}
15469
nipkow
parents:
diff changeset
    80
15476
b8cb20cc0c0b fixed bugs
nipkow
parents: 15469
diff changeset
    81
end
b8cb20cc0c0b fixed bugs
nipkow
parents: 15469
diff changeset
    82
(*>*)