src/HOL/Prolog/prolog.ML
changeset 38786 e46e7a9cb622
parent 38557 9926c47ad1a1
child 38795 848be46708dc
--- a/src/HOL/Prolog/prolog.ML	Thu Aug 26 13:44:50 2010 +0200
+++ b/src/HOL/Prolog/prolog.ML	Thu Aug 26 20:51:17 2010 +0200
@@ -12,7 +12,7 @@
 fun isD t = case t of
     Const(@{const_name Trueprop},_)$t     => isD t
   | Const(@{const_name "op &"}  ,_)$l$r     => isD l andalso isD r
-  | Const(@{const_name "op -->"},_)$l$r     => isG l andalso isD r
+  | Const(@{const_name HOL.implies},_)$l$r     => isG l andalso isD r
   | Const(   "==>",_)$l$r     => isG l andalso isD r
   | Const(@{const_name All},_)$Abs(s,_,t) => isD t
   | Const("all",_)$Abs(s,_,t) => isD t
@@ -32,7 +32,7 @@
     Const(@{const_name Trueprop},_)$t     => isG t
   | Const(@{const_name "op &"}  ,_)$l$r     => isG l andalso isG r
   | Const(@{const_name "op |"}  ,_)$l$r     => isG l andalso isG r
-  | Const(@{const_name "op -->"},_)$l$r     => isD l andalso isG r
+  | Const(@{const_name HOL.implies},_)$l$r     => isD l andalso isG r
   | Const(   "==>",_)$l$r     => isD l andalso isG r
   | Const(@{const_name All},_)$Abs(_,_,t) => isG t
   | Const("all",_)$Abs(_,_,t) => isG t
@@ -54,7 +54,7 @@
       _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS
         (read_instantiate ctxt [(("x", 0), "?" ^ (if s="P" then "PP" else s))] spec))
     | _$(Const(@{const_name "op &"},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
-    | _$(Const(@{const_name "op -->"},_)$_$_)     => at(thm RS mp)
+    | _$(Const(@{const_name HOL.implies},_)$_$_)     => at(thm RS mp)
     | _                             => [thm]
 in map zero_var_indexes (at thm) end;
 
@@ -72,7 +72,7 @@
   -- is nice, but cannot instantiate unknowns in the assumptions *)
 fun hyp_resolve_tac i st = let
         fun ap (Const(@{const_name All},_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
-        |   ap (Const(@{const_name "op -->"},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
+        |   ap (Const(@{const_name HOL.implies},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
         |   ap t                          =                         (0,false,t);
 (*
         fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t