src/HOL/HOL.thy
changeset 1672 2c109cd2fdd0
parent 1370 7361ac9b024d
child 1674 33aff4d854e4
equal deleted inserted replaced
1671:608196238072 1672:2c109cd2fdd0
   141   Inv_def       "Inv(f::'a=>'b)  == (% y. @x. f(x)=y)"
   141   Inv_def       "Inv(f::'a=>'b)  == (% y. @x. f(x)=y)"
   142   o_def         "(f::'b=>'c) o g == (%(x::'a). f(g(x)))"
   142   o_def         "(f::'b=>'c) o g == (%(x::'a). f(g(x)))"
   143   if_def        "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)"
   143   if_def        "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)"
   144 
   144 
   145 (* start 8bit 1 *)
   145 (* start 8bit 1 *)
       
   146 types
       
   147 ('a, 'b) "ë"          (infixr 5)
       
   148 
       
   149 syntax
       
   150   "Ú"		:: "['a::{}, 'a] => prop"       ("(_ Ú/ _)"         [3, 2] 2)
       
   151   "êë"		:: "[prop, prop] => prop"	("(_ êë/ _)"        [2, 1] 1)
       
   152   "Gbigimpl"	:: "[asms, prop] => prop"	("((3Ë _ Ì) êë/ _)" [0, 1] 1)
       
   153   "Gall"	:: "('a => prop) => prop"	(binder "Ä"                0)
       
   154 
       
   155   "Glam"         :: "[idts, 'a::logic] => 'b::logic"  ("(3³_./ _)" 10)
       
   156   "Û"           :: "['a, 'a] => bool"                 (infixl 50)
       
   157   "Gnot"        :: "bool => bool"                     ("¿ _" [40] 40)
       
   158   "GEps"        :: "[pttrn, bool] => 'a"               ("(3®_./ _)" 10)
       
   159   "GAll"        :: "[idts, bool] => bool"             ("(3Â_./ _)" 10)
       
   160   "GEx"         :: "[idts, bool] => bool"             ("(3Ã_./ _)" 10)
       
   161   "GEx1"        :: "[idts, bool] => bool"             ("(3Ã!_./ _)" 10)
       
   162   "À"           :: "[bool, bool] => bool"             (infixr 35)
       
   163   "Á"           :: "[bool, bool] => bool"             (infixr 30)
       
   164   "çè"          :: "[bool, bool] => bool"             (infixr 25)
       
   165 
       
   166 translations
       
   167 (type)  "x ë y"	== (type) "x => y" 
       
   168 
       
   169 (*  "³x.t"	=> "%x.t" *)
       
   170 
       
   171   "x Û y"	== "x ~= y"
       
   172   "¿ y"		== "~y"
       
   173   "®x.P"	== "@x.P"
       
   174   "Âx.P"	== "! x. P"
       
   175   "Ãx.P"	== "? x. P"
       
   176   "Ã!x.P"	== "?! x. P"
       
   177   "x À y"	== "x & y"
       
   178   "x Á y"	== "x | y"
       
   179   "x çè y"	== "x --> y"
   146 (* end 8bit 1 *)
   180 (* end 8bit 1 *)
   147 
   181 
   148 end
   182 end
   149 
   183 
   150 ML
   184 ML
   165 
   199 
   166 val print_ast_translation =
   200 val print_ast_translation =
   167   map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")];
   201   map alt_ast_tr' [("! ", "*All"), ("? ", "*Ex"), ("?! ", "*Ex1")];
   168 
   202 
   169 (* start 8bit 2 *)
   203 (* start 8bit 2 *)
       
   204 local open Ast;
       
   205 fun bigimpl_ast_tr (*"Gbigimpl"*) [asms, concl] =
       
   206       fold_ast_p "êë" (unfold_ast "_asms" asms, concl)
       
   207   | bigimpl_ast_tr (*"Gbigimpl"*) asts = raise_ast "bigimpl_ast_tr" asts;
       
   208 fun impl_ast_tr' (*"êë"*) asts =
       
   209   (case unfold_ast_p "êë" (Appl (Constant "êë" :: asts)) of
       
   210     (asms as _ :: _ :: _, concl)
       
   211       => Appl [Constant "Gbigimpl", fold_ast "_asms" asms, concl]
       
   212   | _ => raise Match);
       
   213 
       
   214 in
       
   215 val parse_ast_translation = ("Gbigimpl", bigimpl_ast_tr)::
       
   216 				("Glam",    fn asts => Appl (Constant "_abs" :: asts))::
       
   217 				parse_ast_translation;
       
   218 
       
   219 val print_ast_translation = ("êë", impl_ast_tr')::
       
   220 				("_lambda", fn asts => Appl (Constant "Glam" :: asts)) ::	
       
   221 				print_ast_translation;
       
   222 
       
   223 end;
       
   224 
       
   225 local open Syntax in
       
   226 val thy = thy 
       
   227 |> add_trrules_i 
       
   228 [mk_appl (Constant "Ú" ) [Variable "P", Variable "Q"] <-> 
       
   229  mk_appl (Constant "==") [Variable "P", Variable "Q"],
       
   230  mk_appl (Constant "êë" ) [Variable "P", Variable "Q"] <-> 
       
   231  mk_appl (Constant "==>") [Variable "P", Variable "Q"],
       
   232  (Constant "Ä" ) <->  (Constant "!!")]
       
   233 end;
   170 (* end 8bit 2 *)
   234 (* end 8bit 2 *)
   171 
   235 
   172 
   236 
   173 
   237