| author | wenzelm | 
| Wed, 27 Mar 2013 14:08:03 +0100 | |
| changeset 51550 | cec08df2c030 | 
| parent 42297 | 140f283266b7 | 
| child 55038 | f2179be64805 | 
| permissions | -rw-r--r-- | 
| 15469 | 1 | (* Title: HOL/Library/OptionalSugar.thy | 
| 2 | Author: Gerwin Klain, Tobias Nipkow | |
| 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: 
30503diff
changeset | 7 | imports Complex_Main LaTeXsugar | 
| 15469 | 8 | begin | 
| 9 | ||
| 30474 | 10 | (* hiding set *) | 
| 22835 | 11 | translations | 
| 30404 
d03dd6301678
more robust treatment of (authentic) consts within translations;
 wenzelm parents: 
29494diff
changeset | 12 | "xs" <= "CONST set xs" | 
| 22835 | 13 | |
| 30474 | 14 | (* hiding numeric conversions - embeddings only *) | 
| 15 | translations | |
| 16 | "n" <= "CONST of_nat n" | |
| 17 | "n" <= "CONST int n" | |
| 35115 | 18 | "n" <= "CONST real n" | 
| 30474 | 19 | "n" <= "CONST real_of_nat n" | 
| 20 | "n" <= "CONST real_of_int n" | |
| 30502 | 21 | "n" <= "CONST of_real n" | 
| 22 | "n" <= "CONST complex_of_real n" | |
| 30474 | 23 | |
| 15469 | 24 | (* append *) | 
| 25 | syntax (latex output) | |
| 35115 | 26 | "_appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^raw:\isacharat>" 65) | 
| 15469 | 27 | translations | 
| 35115 | 28 | "_appendL xs ys" <= "xs @ ys" | 
| 29 | "_appendL (_appendL xs ys) zs" <= "_appendL xs (_appendL ys zs)" | |
| 15469 | 30 | |
| 31 | ||
| 32891 | 32 | (* deprecated, use thm with style instead, will be removed *) | 
| 15469 | 33 | (* aligning equations *) | 
| 21210 | 34 | notation (tab output) | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
35250diff
changeset | 35 |   "HOL.eq"  ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
 | 
| 19674 | 36 |   "=="  ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")
 | 
| 15469 | 37 | |
| 38 | (* Let *) | |
| 39 | translations | |
| 35250 | 40 | "_bind (p, CONST DUMMY) e" <= "_bind p (CONST fst e)" | 
| 41 | "_bind (CONST DUMMY, p) e" <= "_bind p (CONST snd e)" | |
| 15469 | 42 | |
| 43 | "_tuple_args x (_tuple_args y z)" == | |
| 44 | "_tuple_args x (_tuple_arg (_tuple y z))" | |
| 45 | ||
| 35250 | 46 | "_bind (CONST Some p) e" <= "_bind p (CONST the e)" | 
| 47 | "_bind (p # CONST DUMMY) e" <= "_bind p (CONST hd e)" | |
| 48 | "_bind (CONST DUMMY # p) e" <= "_bind p (CONST tl e)" | |
| 15469 | 49 | |
| 29494 | 50 | (* type constraints with spacing *) | 
| 51 | setup {*
 | |
| 52 | let | |
| 33384 | 53 | val typ = Simple_Syntax.read_typ; | 
| 29494 | 54 | in | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 55 | Sign.del_modesyntax_i (Symbol.xsymbolsN, false) | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 56 |    [("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 57 |     ("_constrain", typ "prop' => type => prop'", Mixfix ("_\<Colon>_", [4, 0], 3))] #>
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 58 | Sign.add_modesyntax_i (Symbol.xsymbolsN, false) | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 59 |    [("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 60 |     ("_constrain", typ "prop' => type => prop'", Mixfix ("_ \<Colon> _", [4, 0], 3))]
 | 
| 29494 | 61 | end | 
| 62 | *} | |
| 63 | ||
| 64 | (* sorts as intersections *) | |
| 65 | setup {*
 | |
| 66 | let | |
| 67 |   val sortT = Type ("sort", []); (*FIXME*)
 | |
| 68 |   val classesT = Type ("classes", []); (*FIXME*)
 | |
| 69 | in | |
| 70 | Sign.add_modesyntax_i (Symbol.xsymbolsN, false) [ | |
| 42297 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 wenzelm parents: 
42293diff
changeset | 71 |     ("_topsort", sortT, Mixfix ("\<top>", [], 1000)),
 | 
| 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 wenzelm parents: 
42293diff
changeset | 72 |     ("_sort", classesT --> sortT, Mixfix ("'(_')", [], 1000)),
 | 
| 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 wenzelm parents: 
42293diff
changeset | 73 |     ("_classes", Lexicon.idT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], 1000)),
 | 
| 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 wenzelm parents: 
42293diff
changeset | 74 |     ("_classes", Lexicon.longidT --> classesT --> classesT, Mixfix ("_ \<inter> _", [], 1000))
 | 
| 29494 | 75 | ] | 
| 76 | end | |
| 77 | *} | |
| 15469 | 78 | |
| 15476 | 79 | end | 
| 80 | (*>*) |