--- a/NEWS Tue Mar 03 19:08:04 2015 +0100
+++ b/NEWS Wed Mar 04 19:53:18 2015 +0100
@@ -258,6 +258,10 @@
*** ML ***
+* Elementary operations in module Thm are no longer pervasive.
+INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
+Thm.term_of etc.
+
* Former combinators NAMED_CRITICAL and CRITICAL for central critical
sections have been discontinued, in favour of the more elementary
Multithreading.synchronized and its high-level derivative
--- a/src/CCL/Wfd.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/CCL/Wfd.thy Wed Mar 04 19:53:18 2015 +0100
@@ -423,9 +423,9 @@
| get_bno l n (Bound m) = (m-length(l),n)
(* Not a great way of identifying induction hypothesis! *)
-fun could_IH x = Term.could_unify(x,hd (prems_of @{thm rcallT})) orelse
- Term.could_unify(x,hd (prems_of @{thm rcall2T})) orelse
- Term.could_unify(x,hd (prems_of @{thm rcall3T}))
+fun could_IH x = Term.could_unify(x,hd (Thm.prems_of @{thm rcallT})) orelse
+ Term.could_unify(x,hd (Thm.prems_of @{thm rcall2T})) orelse
+ Term.could_unify(x,hd (Thm.prems_of @{thm rcall3T}))
fun IHinst tac rls = SUBGOAL (fn (Bi,i) =>
let val bvs = bvars Bi []
--- a/src/Doc/Isar_Ref/Generic.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Doc/Isar_Ref/Generic.thy Wed Mar 04 19:53:18 2015 +0100
@@ -999,7 +999,7 @@
simproc_setup unit ("x::unit") = \<open>
fn _ => fn _ => fn ct =>
- if HOLogic.is_unit (term_of ct) then NONE
+ if HOLogic.is_unit (Thm.term_of ct) then NONE
else SOME (mk_meta_eq @{thm unit_eq})
\<close>
--- a/src/FOL/simpdata.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/FOL/simpdata.ML Wed Mar 04 19:53:18 2015 +0100
@@ -7,18 +7,19 @@
(*Make meta-equalities. The operator below is Trueprop*)
-fun mk_meta_eq th = case concl_of th of
+fun mk_meta_eq th =
+ (case Thm.concl_of th of
_ $ (Const(@{const_name eq},_)$_$_) => th RS @{thm eq_reflection}
| _ $ (Const(@{const_name iff},_)$_$_) => th RS @{thm iff_reflection}
- | _ =>
- error("conclusion must be a =-equality or <->");;
+ | _ => error "conclusion must be a =-equality or <->");
-fun mk_eq th = case concl_of th of
+fun mk_eq th =
+ (case Thm.concl_of th of
Const(@{const_name Pure.eq},_)$_$_ => th
| _ $ (Const(@{const_name eq},_)$_$_) => mk_meta_eq th
| _ $ (Const(@{const_name iff},_)$_$_) => mk_meta_eq th
| _ $ (Const(@{const_name Not},_)$_) => th RS @{thm iff_reflection_F}
- | _ => th RS @{thm iff_reflection_T};
+ | _ => th RS @{thm iff_reflection_T});
(*Replace premises x=y, X<->Y by X==Y*)
fun mk_meta_prems ctxt =
@@ -36,16 +37,17 @@
(@{const_name All}, [@{thm spec}]), (@{const_name True}, []), (@{const_name False}, [])];
fun mk_atomize pairs =
- let fun atoms th =
- (case concl_of th of
- Const(@{const_name Trueprop},_) $ p =>
- (case head_of p of
- Const(a,_) =>
- (case AList.lookup (op =) pairs a of
- SOME(rls) => maps atoms ([th] RL rls)
- | NONE => [th])
- | _ => [th])
- | _ => [th])
+ let
+ fun atoms th =
+ (case Thm.concl_of th of
+ Const(@{const_name Trueprop},_) $ p =>
+ (case head_of p of
+ Const(a,_) =>
+ (case AList.lookup (op =) pairs a of
+ SOME(rls) => maps atoms ([th] RL rls)
+ | NONE => [th])
+ | _ => [th])
+ | _ => [th])
in atoms end;
fun mksimps pairs (_: Proof.context) = map mk_eq o mk_atomize pairs o gen_all;
--- a/src/FOLP/simp.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/FOLP/simp.ML Wed Mar 04 19:53:18 2015 +0100
@@ -67,7 +67,7 @@
(*insert a thm in a discrimination net by its lhs*)
fun lhs_insert_thm th net =
- Net.insert_term eq_brl (lhs_of (concl_of th), (false,th)) net
+ Net.insert_term eq_brl (lhs_of (Thm.concl_of th), (false,th)) net
handle Net.INSERT => net;
(*match subgoal i against possible theorems in the net.
@@ -83,7 +83,7 @@
biresolve0_tac (Net.unify_term net
(lhs_of (Logic.strip_assums_concl prem))) i);
-fun nth_subgoal i thm = nth (prems_of thm) (i - 1);
+fun nth_subgoal i thm = nth (Thm.prems_of thm) (i - 1);
fun goal_concl i thm = Logic.strip_assums_concl (nth_subgoal i thm);
@@ -111,7 +111,7 @@
(*Get the norm constants from norm_thms*)
val norms =
let fun norm thm =
- case lhs_of(concl_of thm) of
+ case lhs_of (Thm.concl_of thm) of
Const(n,_)$_ => n
| _ => error "No constant in lhs of a norm_thm"
in map norm normE_thms end;
@@ -145,7 +145,7 @@
(*get name of the constant from conclusion of a congruence rule*)
fun cong_const cong =
- case head_of (lhs_of (concl_of cong)) of
+ case head_of (lhs_of (Thm.concl_of cong)) of
Const(c,_) => c
| _ => "" (*a placeholder distinct from const names*);
@@ -189,10 +189,10 @@
fun add_norms(congs,ccs,new_asms) thm =
let val thm' = mk_trans2 thm;
(* thm': [?z -> l; Prems; r -> ?t] ==> ?z -> ?t *)
- val nops = nprems_of thm'
+ val nops = Thm.nprems_of thm'
val lhs = rhs_of_eq 1 thm'
val rhs = lhs_of_eq nops thm'
- val asms = tl(rev(tl(prems_of thm')))
+ val asms = tl(rev(tl(Thm.prems_of thm')))
val hvars = fold_rev (add_hidden_vars ccs) (lhs::rhs::asms) []
val hvars = add_new_asm_vars new_asms (rhs,hvars)
fun it_asms asm hvars =
@@ -216,7 +216,7 @@
fun add_norm_tags congs =
let val ccs = map cong_const congs
val new_asms = filter (exists not o #2)
- (ccs ~~ (map (map atomic o prems_of) congs));
+ (ccs ~~ (map (map atomic o Thm.prems_of) congs));
in add_norms(congs,ccs,new_asms) end;
fun normed_rews congs =
@@ -252,7 +252,7 @@
(*insert a thm in a thm net*)
fun insert_thm_warn th net =
- Net.insert_term Thm.eq_thm_prop (concl_of th, th) net
+ Net.insert_term Thm.eq_thm_prop (Thm.concl_of th, th) net
handle Net.INSERT =>
(writeln ("Duplicate rewrite or congruence rule:\n" ^
Display.string_of_thm_without_context th); net);
@@ -278,7 +278,7 @@
(*delete a thm from a thm net*)
fun delete_thm_warn th net =
- Net.delete_term Thm.eq_thm_prop (concl_of th, th) net
+ Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net
handle Net.DELETE =>
(writeln ("No such rewrite or congruence rule:\n" ^
Display.string_of_thm_without_context th); net);
@@ -348,7 +348,7 @@
| seq_try([],_) thm = no_tac thm
and try_rew thm = (seq_try(case_rews,case_consts) ORELSE one_subt) thm
and one_subt thm =
- let val test = has_fewer_prems (nprems_of thm + 1)
+ let val test = has_fewer_prems (Thm.nprems_of thm + 1)
fun loop thm =
COND test no_tac
((try_rew THEN DEPTH_FIRST test (refl_tac i))
@@ -424,8 +424,8 @@
if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs)
else case Seq.pull(cong_tac i thm) of
SOME(thm',_) =>
- let val ps = prems_of thm
- and ps' = prems_of thm';
+ let val ps = Thm.prems_of thm
+ and ps' = Thm.prems_of thm';
val n = length(ps')-length(ps);
val a = length(Logic.strip_assums_hyp(nth ps (i - 1)))
val l = map (length o Logic.strip_assums_hyp) (take n (drop (i-1) ps'));
@@ -436,7 +436,7 @@
are represented by rules, generalized over their parameters*)
fun add_asms(ss,thm,a,anet,ats,cs) =
let val As = strip_varify(nth_subgoal i thm, a, []);
- val thms = map (Thm.trivial o cterm_of(Thm.theory_of_thm thm)) As;
+ val thms = map (Thm.trivial o Thm.cterm_of(Thm.theory_of_thm thm)) As;
val new_rws = maps mk_rew_rules thms;
val rwrls = map mk_trans (maps mk_rew_rules thms);
val anet' = fold_rev lhs_insert_thm rwrls anet;
@@ -449,7 +449,7 @@
fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of
SOME(thm',seq') =>
- let val n = (nprems_of thm') - (nprems_of thm)
+ let val n = (Thm.nprems_of thm') - (Thm.nprems_of thm)
in pr_rew(i,n,thm,thm',more);
if n=0 then (SIMP_LHS::ss, thm', anet, ats, cs)
else ((replicate n PROVE) @ (POP_CS::SIMP_LHS::ss),
@@ -499,7 +499,7 @@
let val cong_tac = net_tac cong_net
in fn i =>
(fn thm =>
- if i <= 0 orelse nprems_of thm < i then Seq.empty
+ if i <= 0 orelse Thm.nprems_of thm < i then Seq.empty
else Seq.single(execute(ss,fl,auto_tac,cong_tac,simp_net,i,thm)))
THEN TRY(auto_tac i)
end;
@@ -547,8 +547,8 @@
fun find_subst sg T =
let fun find (thm::thms) =
- let val (Const(_,cT), va, vb) = dest_red(hd(prems_of thm));
- val [P] = subtract (op =) [va, vb] (Misc_Legacy.add_term_vars (concl_of thm, []));
+ let val (Const(_,cT), va, vb) = dest_red(hd(Thm.prems_of thm));
+ val [P] = subtract (op =) [va, vb] (Misc_Legacy.add_term_vars (Thm.concl_of thm, []));
val eqT::_ = binder_types cT
in if Sign.typ_instance sg (T,eqT) then SOME(thm,va,vb,P)
else find thms
@@ -559,12 +559,12 @@
fun mk_cong sg (f,aTs,rT) (refl,eq) =
let val k = length aTs;
fun ri((subst,va as Var(_,Ta),vb as Var(_,Tb),P),i,si,T,yik) =
- let val ca = cterm_of sg va
- and cx = cterm_of sg (eta_Var(("X"^si,0),T))
- val cb = cterm_of sg vb
- and cy = cterm_of sg (eta_Var(("Y"^si,0),T))
- val cP = cterm_of sg P
- and cp = cterm_of sg (Pinst(f,rT,eq,k,i,T,yik,aTs))
+ let val ca = Thm.cterm_of sg va
+ and cx = Thm.cterm_of sg (eta_Var(("X"^si,0),T))
+ val cb = Thm.cterm_of sg vb
+ and cy = Thm.cterm_of sg (eta_Var(("Y"^si,0),T))
+ val cP = Thm.cterm_of sg P
+ and cp = Thm.cterm_of sg (Pinst(f,rT,eq,k,i,T,yik,aTs))
in cterm_instantiate [(ca,cx),(cb,cy),(cP,cp)] subst end;
fun mk(c,T::Ts,i,yik) =
let val si = radixstring(26,"a",i)
@@ -579,7 +579,7 @@
fun mk_cong_type sg (f,T) =
let val (aTs,rT) = strip_type T;
fun find_refl(r::rs) =
- let val (Const(eq,eqT),_,_) = dest_red(concl_of r)
+ let val (Const(eq,eqT),_,_) = dest_red(Thm.concl_of r)
in if Sign.typ_instance sg (rT, hd(binder_types eqT))
then SOME(r,(eq,body_type eqT)) else find_refl rs
end
--- a/src/FOLP/simpdata.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/FOLP/simpdata.ML Wed Mar 04 19:53:18 2015 +0100
@@ -16,7 +16,7 @@
(* Conversion into rewrite rules *)
-fun mk_eq th = case concl_of th of
+fun mk_eq th = case Thm.concl_of th of
_ $ (Const (@{const_name iff}, _) $ _ $ _) $ _ => th
| _ $ (Const (@{const_name eq}, _) $ _ $ _) $ _ => th
| _ $ (Const (@{const_name Not}, _) $ _) $ _ => th RS @{thm not_P_imp_P_iff_F}
@@ -32,7 +32,7 @@
fun mk_atomize pairs =
let fun atoms th =
- (case concl_of th of
+ (case Thm.concl_of th of
Const ("Trueprop", _) $ p =>
(case head_of p of
Const(a,_) =>
--- a/src/HOL/Decision_Procs/Approximation.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy Wed Mar 04 19:53:18 2015 +0100
@@ -3597,11 +3597,11 @@
| variable_of_bound t = raise TERM ("variable_of_bound", [t])
val variable_bounds
- = map (` (variable_of_bound o prop_of)) prems
+ = map (`(variable_of_bound o Thm.prop_of)) prems
fun add_deps (name, bnds)
= Graph.add_deps_acyclic (name,
- remove (op =) name (Term.add_free_names (prop_of bnds) []))
+ remove (op =) name (Term.add_free_names (Thm.prop_of bnds) []))
val order = Graph.empty
|> fold Graph.new_node variable_bounds
@@ -3634,7 +3634,7 @@
= case AList.lookup (op =) splitting name
of SOME s => HOLogic.mk_number @{typ nat} s
| NONE => @{term "0 :: nat"}
- val vs = nth (prems_of st) (i - 1)
+ val vs = nth (Thm.prems_of st) (i - 1)
|> Logic.strip_imp_concl
|> HOLogic.dest_Trueprop
|> Term.strip_comb |> snd |> List.last
@@ -3659,7 +3659,9 @@
THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st
end
- | SOME t => if length vs <> 1 then raise (TERM ("More than one variable used for taylor series expansion", [prop_of st]))
+ | SOME t =>
+ if length vs <> 1
+ then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st]))
else let
val t = t
|> HOLogic.mk_number @{typ nat}
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Wed Mar 04 19:53:18 2015 +0100
@@ -558,20 +558,20 @@
let
val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
fun h x t =
- case term_of t of
+ case Thm.term_of t of
Const(@{const_name HOL.eq}, _)$y$z =>
- if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq
else Ferrante_Rackoff_Data.Nox
| @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
- if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq
else Ferrante_Rackoff_Data.Nox
| b$y$z => if Term.could_unify (b, lt) then
- if term_of x aconv y then Ferrante_Rackoff_Data.Lt
- else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt
+ else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt
else Ferrante_Rackoff_Data.Nox
else if Term.could_unify (b, le) then
- if term_of x aconv y then Ferrante_Rackoff_Data.Le
- else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le
+ else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge
else Ferrante_Rackoff_Data.Nox
else Ferrante_Rackoff_Data.Nox
| _ => Ferrante_Rackoff_Data.Nox
@@ -709,7 +709,7 @@
if h aconvc y then false else if h aconvc x then true else earlier t x y;
fun dest_frac ct =
- case term_of ct of
+ case Thm.term_of ct of
Const (@{const_name Fields.divide},_) $ a $ b=>
Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
| Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
@@ -724,21 +724,21 @@
(Numeral.mk_cnumber cT b)
end
-fun whatis x ct = case term_of ct of
+fun whatis x ct = case Thm.term_of ct of
Const(@{const_name Groups.plus}, _)$(Const(@{const_name Groups.times},_)$_$y)$_ =>
- if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
+ if y aconv Thm.term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct])
else ("Nox",[])
| Const(@{const_name Groups.plus}, _)$y$_ =>
- if y aconv term_of x then ("x+t",[Thm.dest_arg ct])
+ if y aconv Thm.term_of x then ("x+t",[Thm.dest_arg ct])
else ("Nox",[])
| Const(@{const_name Groups.times}, _)$_$y =>
- if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct])
+ if y aconv Thm.term_of x then ("c*x",[Thm.dest_arg1 ct])
else ("Nox",[])
-| t => if t aconv term_of x then ("x",[]) else ("Nox",[]);
+| t => if t aconv Thm.term_of x then ("x",[]) else ("Nox",[]);
fun xnormalize_conv ctxt [] ct = Thm.reflexive ct
| xnormalize_conv ctxt (vs as (x::_)) ct =
- case term_of ct of
+ case Thm.term_of ct of
Const(@{const_name Orderings.less},_)$_$Const(@{const_name Groups.zero},_) =>
(case whatis x (Thm.dest_arg1 ct) of
("c*x+t",[c,t]) =>
@@ -752,14 +752,14 @@
(if neg then Thm.apply (Thm.apply clt c) cz
else Thm.apply (Thm.apply clt cz) c))
val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
- val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x,t])
+ val th = Thm.implies_elim (instantiate' [SOME (Thm.ctyp_of_term x)] (map SOME [c,x,t])
(if neg then @{thm neg_prod_sum_lt} else @{thm pos_prod_sum_lt})) cth
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
in rth end
| ("x+t",[t]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"}
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
@@ -775,7 +775,7 @@
(if neg then Thm.apply (Thm.apply clt c) cz
else Thm.apply (Thm.apply clt cz) c))
val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
- val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+ val th = Thm.implies_elim (instantiate' [SOME (Thm.ctyp_of_term x)] (map SOME [c,x])
(if neg then @{thm neg_prod_lt} else @{thm pos_prod_lt})) cth
val rth = th
in rth end
@@ -786,7 +786,7 @@
(case whatis x (Thm.dest_arg1 ct) of
("c*x+t",[c,t]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val cr = dest_frac c
val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
val cz = Thm.dest_arg ct
@@ -803,14 +803,14 @@
in rth end
| ("x+t",[t]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"}
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
in rth end
| ("c*x",[c]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val cr = dest_frac c
val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"}
val cz = Thm.dest_arg ct
@@ -820,7 +820,7 @@
(if neg then Thm.apply (Thm.apply clt c) cz
else Thm.apply (Thm.apply clt cz) c))
val cth = Thm.equal_elim (Thm.symmetric cthp) TrueI
- val th = Thm.implies_elim (instantiate' [SOME (ctyp_of_term x)] (map SOME [c,x])
+ val th = Thm.implies_elim (instantiate' [SOME (Thm.ctyp_of_term x)] (map SOME [c,x])
(if neg then @{thm neg_prod_le} else @{thm pos_prod_le})) cth
val rth = th
in rth end
@@ -830,7 +830,7 @@
(case whatis x (Thm.dest_arg1 ct) of
("c*x+t",[c,t]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val cr = dest_frac c
val ceq = Thm.dest_fun2 ct
val cz = Thm.dest_arg ct
@@ -845,14 +845,14 @@
in rth end
| ("x+t",[t]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"}
val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv
(Semiring_Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th
in rth end
| ("c*x",[c]) =>
let
- val T = ctyp_of_term x
+ val T = Thm.ctyp_of_term x
val cr = dest_frac c
val ceq = Thm.dest_fun2 ct
val cz = Thm.dest_arg ct
@@ -871,10 +871,10 @@
val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"}
val ss = simpset_of @{context}
in
-fun field_isolate_conv phi ctxt vs ct = case term_of ct of
+fun field_isolate_conv phi ctxt vs ct = case Thm.term_of ct of
Const(@{const_name Orderings.less},_)$a$b =>
let val (ca,cb) = Thm.dest_binop ct
- val T = ctyp_of_term ca
+ val T = Thm.ctyp_of_term ca
val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
@@ -883,7 +883,7 @@
in rth end
| Const(@{const_name Orderings.less_eq},_)$a$b =>
let val (ca,cb) = Thm.dest_binop ct
- val T = ctyp_of_term ca
+ val T = Thm.ctyp_of_term ca
val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
@@ -893,7 +893,7 @@
| Const(@{const_name HOL.eq},_)$a$b =>
let val (ca,cb) = Thm.dest_binop ct
- val T = ctyp_of_term ca
+ val T = Thm.ctyp_of_term ca
val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
val nth = Conv.fconv_rule
(Conv.arg_conv (Conv.arg1_conv
@@ -907,19 +907,21 @@
fun classfield_whatis phi =
let
fun h x t =
- case term_of t of
- Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
- else Ferrante_Rackoff_Data.Nox
- | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
- else Ferrante_Rackoff_Data.Nox
+ case Thm.term_of t of
+ Const(@{const_name HOL.eq}, _)$y$z =>
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Eq
+ else Ferrante_Rackoff_Data.Nox
+ | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) =>
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.NEq
+ else Ferrante_Rackoff_Data.Nox
| Const(@{const_name Orderings.less},_)$y$z =>
- if term_of x aconv y then Ferrante_Rackoff_Data.Lt
- else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
- else Ferrante_Rackoff_Data.Nox
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Lt
+ else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Gt
+ else Ferrante_Rackoff_Data.Nox
| Const (@{const_name Orderings.less_eq},_)$y$z =>
- if term_of x aconv y then Ferrante_Rackoff_Data.Le
- else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
- else Ferrante_Rackoff_Data.Nox
+ if Thm.term_of x aconv y then Ferrante_Rackoff_Data.Le
+ else if Thm.term_of x aconv z then Ferrante_Rackoff_Data.Ge
+ else Ferrante_Rackoff_Data.Nox
| _ => Ferrante_Rackoff_Data.Nox
in h end;
fun class_field_ss phi =
--- a/src/HOL/Decision_Procs/MIR.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Wed Mar 04 19:53:18 2015 +0100
@@ -5655,7 +5655,7 @@
val vs = map_index swap fs;
val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe};
val t' = (term_of_fm vs o qe o fm_of_term vs) t;
- in (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end
+ in (Thm.cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end
end;
*}
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Mar 04 19:53:18 2015 +0100
@@ -4156,7 +4156,7 @@
fn (ctxt, alternative, ty, ps, ct) =>
Proof_Context.cterm_of ctxt
- (frpar_procedure alternative ty ps (term_of ct))
+ (frpar_procedure alternative ty ps (Thm.term_of ct))
end
*}
--- a/src/HOL/Decision_Procs/approximation.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/approximation.ML Wed Mar 04 19:53:18 2015 +0100
@@ -95,7 +95,7 @@
Proof_Context.cterm_of ctxt term
|> Goal.init
|> SINGLE tactic
- |> the |> prems_of |> hd
+ |> the |> Thm.prems_of |> hd
fun prepare_form ctxt term = apply_tactic ctxt term (
REPEAT (FIRST' [etac @{thm intervalE}, etac @{thm meta_eqE}, rtac @{thm impI}] 1)
@@ -122,7 +122,7 @@
fun approx_arith prec ctxt t = realify t
|> Proof_Context.cterm_of ctxt
|> Reification.conv ctxt form_equations
- |> prop_of
+ |> Thm.prop_of
|> Logic.dest_equals |> snd
|> dest_interpret |> fst
|> mk_approx' prec
--- a/src/HOL/Decision_Procs/approximation_generator.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/approximation_generator.ML Wed Mar 04 19:53:18 2015 +0100
@@ -117,7 +117,7 @@
}
fun rewrite_with ctxt thms = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps thms)
-fun conv_term thy conv r = cterm_of thy r |> conv |> Thm.prop_of |> Logic.dest_equals |> snd
+fun conv_term thy conv r = Thm.cterm_of thy r |> conv |> Thm.prop_of |> Logic.dest_equals |> snd
fun approx_random ctxt prec eps frees e xs genuine_only size seed =
let
--- a/src/HOL/Decision_Procs/commutative_ring_tac.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Wed Mar 04 19:53:18 2015 +0100
@@ -78,10 +78,10 @@
let
val thy = Proof_Context.theory_of ctxt;
val fs = Misc_Legacy.term_frees eq;
- val cvs = cterm_of thy (HOLogic.mk_list T fs);
- val clhs = cterm_of thy (reif_polex T fs lhs);
- val crhs = cterm_of thy (reif_polex T fs rhs);
- val ca = ctyp_of thy T;
+ val cvs = Thm.cterm_of thy (HOLogic.mk_list T fs);
+ val clhs = Thm.cterm_of thy (reif_polex T fs lhs);
+ val crhs = Thm.cterm_of thy (reif_polex T fs rhs);
+ val ca = Thm.ctyp_of thy T;
in (ca, cvs, clhs, crhs) end
else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort ctxt cr_sort)
| reif_eq _ _ = error "reif_eq: not an equation";
--- a/src/HOL/Decision_Procs/cooper_tac.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/cooper_tac.ML Wed Mar 04 19:53:18 2015 +0100
@@ -73,7 +73,7 @@
|> fold Simplifier.add_cong @{thms conj_le_cong imp_le_cong}
(* simp rules for elimination of abs *)
val simpset3 = put_simpset HOL_basic_ss ctxt |> Splitter.add_split @{thm abs_split}
- val ct = cterm_of thy (HOLogic.mk_Trueprop t)
+ val ct = Thm.cterm_of thy (HOLogic.mk_Trueprop t)
(* Theorem for the nat --> int transformation *)
val pre_thm = Seq.hd (EVERY
[simp_tac mod_div_simpset 1, simp_tac simpset0 1,
@@ -83,10 +83,10 @@
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
(* The result of the quantifier elimination *)
val (th, tac) =
- (case (prop_of pre_thm) of
+ (case Thm.prop_of pre_thm of
Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
let
- val pth = linzqe_oracle (cterm_of thy (Envir.eta_long [] t1))
+ val pth = linzqe_oracle (Thm.cterm_of thy (Envir.eta_long [] t1))
in
((pth RS iffD2) RS pre_thm,
assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
--- a/src/HOL/Decision_Procs/ferrack_tac.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML Wed Mar 04 19:53:18 2015 +0100
@@ -55,7 +55,7 @@
val (t,np,nh) = prepare_for_linr q g
(* Some simpsets for dealing with mod div abs and nat*)
val simpset0 = put_simpset HOL_basic_ss ctxt addsimps comp_arith
- val ct = cterm_of thy (HOLogic.mk_Trueprop t)
+ val ct = Thm.cterm_of thy (HOLogic.mk_Trueprop t)
(* Theorem for the nat --> int transformation *)
val pre_thm = Seq.hd (EVERY
[simp_tac simpset0 1,
@@ -63,7 +63,7 @@
(Thm.trivial ct))
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
(* The result of the quantifier elimination *)
- val (th, tac) = case prop_of pre_thm of
+ val (th, tac) = case Thm.prop_of pre_thm of
Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
let val pth = linr_oracle (ctxt, Envir.eta_long [] t1)
in
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML Wed Mar 04 19:53:18 2015 +0100
@@ -25,14 +25,14 @@
fun get_p1 th =
funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE)
- (funpow 2 Thm.dest_arg (cprop_of th)) |> Thm.dest_arg
+ (funpow 2 Thm.dest_arg (Thm.cprop_of th)) |> Thm.dest_arg
fun ferrack_conv ctxt
(entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi,
ld = ld, qe = qe, atoms = atoms},
{isolate_conv = icv, whatis = wi, simpset = simpset}):entry) =
let
- fun uset (vars as (x::vs)) p = case term_of p of
+ fun uset (vars as (x::vs)) p = case Thm.term_of p of
Const(@{const_name HOL.conj}, _)$ _ $ _ =>
let
val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
@@ -56,7 +56,7 @@
val ((p1_v,p2_v),(mp1_v,mp2_v)) =
funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE)
- (funpow 4 Thm.dest_arg (cprop_of (hd minf)))
+ (funpow 4 Thm.dest_arg (Thm.cprop_of (hd minf)))
|> Thm.dest_binop |> apply2 Thm.dest_binop |> apfst (apply2 Thm.dest_fun)
fun myfwd (th1, th2, th3, th4, th5) p1 p2
@@ -73,14 +73,14 @@
in (fw true th1 th_1 th_1', fw false th2 th_2 th_2',
fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5')
end
- val U_v = (Thm.dest_arg o Thm.dest_arg o Thm.dest_arg1) (cprop_of qe)
+ val U_v = (Thm.dest_arg o Thm.dest_arg o Thm.dest_arg1) (Thm.cprop_of qe)
fun main vs p =
let
- val ((xn,ce),(x,fm)) = (case term_of p of
+ val ((xn,ce),(x,fm)) = (case Thm.term_of p of
Const(@{const_name Ex},_)$Abs(xn,xT,_) =>
Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn
| _ => raise CTERM ("main QE only treats existential quantifiers!", [p]))
- val cT = ctyp_of_term x
+ val cT = Thm.ctyp_of_term x
val (u,nth) = uset (x::vs) fm |>> distinct (op aconvc)
val nthx = Thm.abstract_rule xn x nth
val q = Thm.rhs_of nth
@@ -97,18 +97,18 @@
val insI2 = instantiate' [SOME cT] [] @{thm "insertI2"}
in
fun provein x S =
- case term_of S of
+ case Thm.term_of S of
Const(@{const_name Orderings.bot}, _) => raise CTERM ("provein : not a member!", [S])
| Const(@{const_name insert}, _) $ y $_ =>
let val (cy,S') = Thm.dest_binop S
- in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
+ in if Thm.term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
(provein x S')
end
end
- val tabU = fold (fn t => fn tab => Termtab.update (term_of t, provein t cU) tab)
+ val tabU = fold (fn t => fn tab => Termtab.update (Thm.term_of t, provein t cU) tab)
u Termtab.empty
- val U = the o Termtab.lookup tabU o term_of
+ val U = the o Termtab.lookup tabU o Thm.term_of
val [minf_conj, minf_disj, minf_eq, minf_neq, minf_lt,
minf_le, minf_gt, minf_ge, minf_P] = minf
val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt,
@@ -121,7 +121,7 @@
ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) ld
fun decomp_mpinf fm =
- case term_of fm of
+ case Thm.term_of fm of
Const(@{const_name HOL.conj},_)$_$_ =>
let val (p,q) = Thm.dest_binop fm
in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj)
@@ -174,7 +174,7 @@
val grab_atom_bop =
let
fun h bounds tm =
- (case term_of tm of
+ (case Thm.term_of tm of
Const (@{const_name HOL.eq}, T) $ _ $ _ =>
if domain_type T = HOLogic.boolT then find_args bounds tm
else Thm.dest_fun2 tm
--- a/src/HOL/Decision_Procs/langford.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/langford.ML Wed Mar 04 19:53:18 2015 +0100
@@ -14,7 +14,7 @@
val dest_set =
let
fun h acc ct =
- (case term_of ct of
+ (case Thm.term_of ct of
Const (@{const_name Orderings.bot}, _) => acc
| Const (@{const_name insert}, _) $ _ $ t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct));
in h [] end;
@@ -33,7 +33,7 @@
(Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms ball_simps simp_thms})));
fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep =
- (case term_of ep of
+ (case Thm.term_of ep of
Const (@{const_name Ex}, _) $ _ =>
let
val p = Thm.dest_arg ep
@@ -46,13 +46,13 @@
fun proveneF S =
let
val (a, A) = Thm.dest_comb S |>> Thm.dest_arg
- val cT = ctyp_of_term a
+ val cT = Thm.ctyp_of_term a
val ne = instantiate' [SOME cT] [SOME a, SOME A] @{thm insert_not_empty}
val f = prove_finite cT (dest_set S)
in (ne, f) end
val qe =
- (case (term_of L, term_of U) of
+ (case (Thm.term_of L, Thm.term_of U) of
(Const (@{const_name Orderings.bot}, _),_) =>
let val (neU, fU) = proveneF U
in simp_rule ctxt (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end
@@ -70,13 +70,13 @@
val all_conjuncts =
let
fun h acc ct =
- (case term_of ct of
+ (case Thm.term_of ct of
@{term HOL.conj} $ _ $ _ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
| _ => ct :: acc)
in h [] end;
fun conjuncts ct =
- (case term_of ct of
+ (case Thm.term_of ct of
@{term HOL.conj} $ _ $ _ => Thm.dest_arg1 ct :: conjuncts (Thm.dest_arg ct)
| _ => [ct]);
@@ -88,7 +88,7 @@
fun mk_conj_tab th =
let
fun h acc th =
- (case prop_of th of
+ (case Thm.prop_of th of
@{term "Trueprop"} $ (@{term HOL.conj} $ p $ q) =>
h (h acc (th RS conjunct2)) (th RS conjunct1)
| @{term "Trueprop"} $ p => (p, th) :: acc)
@@ -100,7 +100,7 @@
fun prove_conj tab cjs =
(case cjs of
[c] =>
- if is_conj (term_of c)
+ if is_conj (Thm.term_of c)
then prove_conj tab (conjuncts c)
else tab c
| c :: cs => conjI OF [prove_conj tab [c], prove_conj tab cs]);
@@ -108,8 +108,8 @@
fun conj_aci_rule eq =
let
val (l, r) = Thm.dest_equals eq
- fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (term_of c))
- fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (term_of c))
+ fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (Thm.term_of c))
+ fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (Thm.term_of c))
val ll = Thm.dest_arg l
val rr = Thm.dest_arg r
val thl = prove_conj tabl (conjuncts rr) |> Drule.implies_intr_hyps
@@ -118,18 +118,18 @@
in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end;
fun contains x ct =
- member (op aconv) (Misc_Legacy.term_frees (term_of ct)) (term_of x);
+ member (op aconv) (Misc_Legacy.term_frees (Thm.term_of ct)) (Thm.term_of x);
fun is_eqx x eq =
- (case term_of eq of
+ (case Thm.term_of eq of
Const (@{const_name HOL.eq}, _) $ l $ r =>
- l aconv term_of x orelse r aconv term_of x
+ l aconv Thm.term_of x orelse r aconv Thm.term_of x
| _ => false);
local
fun proc ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
Const (@{const_name Ex}, _) $ Abs (xn, _, _) =>
let
val e = Thm.dest_fun ct
@@ -198,7 +198,7 @@
val grab_atom_bop =
let
fun h bounds tm =
- (case term_of tm of
+ (case Thm.term_of tm of
Const (@{const_name HOL.eq}, T) $ _ $ _ =>
if domain_type T = HOLogic.boolT then find_args bounds tm
else Thm.dest_fun2 tm
@@ -231,8 +231,8 @@
fun generalize_tac f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st =>
let
fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "Pure.all"}
- fun gen x t = Thm.apply (all (ctyp_of_term x)) (Thm.lambda x t)
- val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
+ fun gen x t = Thm.apply (all (Thm.ctyp_of_term x)) (Thm.lambda x t)
+ val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (Thm.term_of a, Thm.term_of b)) (f p)
val p' = fold_rev gen ts p
in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end));
@@ -240,7 +240,7 @@
let
val ins = insert (op aconvc)
fun h acc t =
- (case (term_of t) of
+ (case Thm.term_of t of
_ $ _ $ _ =>
if member (op aconvc) ats (Thm.dest_fun2 t)
then ins (Thm.dest_arg t) (ins (Thm.dest_arg1 t) acc)
--- a/src/HOL/Decision_Procs/mir_tac.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Wed Mar 04 19:53:18 2015 +0100
@@ -101,7 +101,7 @@
@{thm "int_0"}, @{thm "int_1"}]
|> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]
(* simp rules for elimination of abs *)
- val ct = cterm_of thy (HOLogic.mk_Trueprop t)
+ val ct = Thm.cterm_of thy (HOLogic.mk_Trueprop t)
(* Theorem for the nat --> int transformation *)
val pre_thm = Seq.hd (EVERY
[simp_tac mod_div_simpset 1, simp_tac simpset0 1,
@@ -110,13 +110,14 @@
(Thm.trivial ct))
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
(* The result of the quantifier elimination *)
- val (th, tac) = case (prop_of pre_thm) of
+ val (th, tac) =
+ case Thm.prop_of pre_thm of
Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
let val pth =
(* If quick_and_dirty then run without proof generation as oracle*)
if Config.get ctxt quick_and_dirty
- then mirfr_oracle (false, cterm_of thy (Envir.eta_long [] t1))
- else mirfr_oracle (true, cterm_of thy (Envir.eta_long [] t1))
+ then mirfr_oracle (false, Thm.cterm_of thy (Envir.eta_long [] t1))
+ else mirfr_oracle (true, Thm.cterm_of thy (Envir.eta_long [] t1))
in
((pth RS iffD2) RS pre_thm,
assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
--- a/src/HOL/Enum.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Enum.thy Wed Mar 04 19:53:18 2015 +0100
@@ -563,12 +563,14 @@
by(cases x) simp
simproc_setup finite_1_eq ("x::finite_1") = {*
- fn _ => fn _ => fn ct => case term_of ct of
- Const (@{const_name a\<^sub>1}, _) => NONE
- | _ => SOME (mk_meta_eq @{thm finite_1_eq})
+ fn _ => fn _ => fn ct =>
+ (case Thm.term_of ct of
+ Const (@{const_name a\<^sub>1}, _) => NONE
+ | _ => SOME (mk_meta_eq @{thm finite_1_eq}))
*}
-instantiation finite_1 :: complete_boolean_algebra begin
+instantiation finite_1 :: complete_boolean_algebra
+begin
definition [simp]: "op - = (\<lambda>_ _. a\<^sub>1)"
definition [simp]: "uminus = (\<lambda>_. a\<^sub>1)"
instance by intro_classes simp_all
--- a/src/HOL/HOL.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/HOL.thy Wed Mar 04 19:53:18 2015 +0100
@@ -1193,14 +1193,14 @@
simproc_setup let_simp ("Let x f") = {*
let
val (f_Let_unfold, x_Let_unfold) =
- let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_unfold}
- in (cterm_of @{theory} f, cterm_of @{theory} x) end
+ let val [(_ $ (f $ x) $ _)] = Thm.prems_of @{thm Let_unfold}
+ in (Thm.cterm_of @{theory} f, Thm.cterm_of @{theory} x) end
val (f_Let_folded, x_Let_folded) =
- let val [(_ $ (f $ x) $ _)] = prems_of @{thm Let_folded}
- in (cterm_of @{theory} f, cterm_of @{theory} x) end;
+ let val [(_ $ (f $ x) $ _)] = Thm.prems_of @{thm Let_folded}
+ in (Thm.cterm_of @{theory} f, Thm.cterm_of @{theory} x) end;
val g_Let_folded =
- let val [(_ $ _ $ (g $ _))] = prems_of @{thm Let_folded}
- in cterm_of @{theory} g end;
+ let val [(_ $ _ $ (g $ _))] = Thm.prems_of @{thm Let_folded}
+ in Thm.cterm_of @{theory} g end;
fun count_loose (Bound i) k = if i >= k then 1 else 0
| count_loose (s $ t) k = count_loose s k + count_loose t k
| count_loose (Abs (_, _, t)) k = count_loose t (k + 1)
@@ -1222,11 +1222,11 @@
else
let
val n = case f of (Abs (x, _, _)) => x | _ => "x";
- val cx = cterm_of thy x;
- val {T = xT, ...} = rep_cterm cx;
- val cf = cterm_of thy f;
+ val cx = Thm.cterm_of thy x;
+ val {T = xT, ...} = Thm.rep_cterm cx;
+ val cf = Thm.cterm_of thy f;
val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx);
- val (_ $ _ $ g) = prop_of fx_g;
+ val (_ $ _ $ g) = Thm.prop_of fx_g;
val g' = abstract_over (x,g);
val abs_g'= Abs (n,xT,g');
in (if (g aconv g')
@@ -1238,10 +1238,10 @@
else if (Envir.beta_eta_contract f) aconv (Envir.beta_eta_contract abs_g') then NONE (*avoid identity conversion*)
else let
val g'x = abs_g'$x;
- val g_g'x = Thm.symmetric (Thm.beta_conversion false (cterm_of thy g'x));
+ val g_g'x = Thm.symmetric (Thm.beta_conversion false (Thm.cterm_of thy g'x));
val rl = cterm_instantiate
- [(f_Let_folded, cterm_of thy f), (x_Let_folded, cx),
- (g_Let_folded, cterm_of thy abs_g')]
+ [(f_Let_folded, Thm.cterm_of thy f), (x_Let_folded, cx),
+ (g_Let_folded, Thm.cterm_of thy abs_g')]
@{thm Let_folded};
in SOME (rl OF [Thm.transitive fx_g g_g'x])
end)
--- a/src/HOL/HOLCF/Cfun.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/HOLCF/Cfun.thy Wed Mar 04 19:53:18 2015 +0100
@@ -144,7 +144,7 @@
let
val dest = Thm.dest_comb;
val f = (snd o dest o snd o dest) ct;
- val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
+ val [T, U] = Thm.dest_ctyp (Thm.ctyp_of_term f);
val tr = instantiate' [SOME T, SOME U] [SOME f]
(mk_meta_eq @{thm Abs_cfun_inverse2});
val rules = Named_Theorems.get ctxt @{named_theorems cont2cont};
--- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Wed Mar 04 19:53:18 2015 +0100
@@ -135,7 +135,7 @@
fun beta_of_def thy def_thm =
let
val (con, lam) =
- Logic.dest_equals (Logic.unvarify_global (concl_of def_thm))
+ Logic.dest_equals (Logic.unvarify_global (Thm.concl_of def_thm))
val (args, rhs) = arglist lam
val lhs = list_ccomb (con, args)
val goal = mk_equals (lhs, rhs)
@@ -211,7 +211,7 @@
val (n1, t1) = args2typ n (snd con)
val (n2, t2) = cons2typ n1 cons
in (n2, mk_ssumT (t1, t2)) end
- val ct = ctyp_of thy (snd (cons2typ 1 spec'))
+ val ct = Thm.ctyp_of thy (snd (cons2typ 1 spec'))
val thm1 = instantiate' [SOME ct] [] @{thm exh_start}
val thm2 = rewrite_rule (Proof_Context.init_global thy)
(map mk_meta_eq @{thms ex_bottom_iffs}) thm1
--- a/src/HOL/HOLCF/Tools/cont_proc.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/HOLCF/Tools/cont_proc.ML Wed Mar 04 19:53:18 2015 +0100
@@ -122,7 +122,7 @@
fun solve_cont ctxt t =
let
val thy = Proof_Context.theory_of ctxt
- val tr = instantiate' [] [SOME (cterm_of thy t)] @{thm Eq_TrueI}
+ val tr = instantiate' [] [SOME (Thm.cterm_of thy t)] @{thm Eq_TrueI}
in Option.map fst (Seq.pull (cont_tac ctxt 1 tr)) end
in
fun cont_proc thy =
--- a/src/HOL/HOLCF/Tools/fixrec.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/HOLCF/Tools/fixrec.ML Wed Mar 04 19:53:18 2015 +0100
@@ -100,7 +100,7 @@
| name_of t = raise TERM ("Fixrec.add_unfold: lhs not a constant", [t])
val lhs_name =
- name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of
+ name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of
in
--- a/src/HOL/Import/import_data.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Import/import_data.ML Wed Mar 04 19:53:18 2015 +0100
@@ -73,7 +73,7 @@
val th = Thm.legacy_freezeT th
val name = case name_opt of
NONE => (fst o dest_Const o fst o HOLogic.dest_eq o
- HOLogic.dest_Trueprop o prop_of) th
+ HOLogic.dest_Trueprop o Thm.prop_of) th
| SOME n => n
val thy' = add_const_map s name thy
in
@@ -85,7 +85,7 @@
fun add_typ_def tyname absname repname th thy =
let
val th = Thm.legacy_freezeT th
- val (l, _) = dest_comb (HOLogic.dest_Trueprop (prop_of th))
+ val (l, _) = dest_comb (HOLogic.dest_Trueprop (Thm.prop_of th))
val (l, abst) = dest_comb l
val (_, rept) = dest_comb l
val (absn, _) = dest_Const abst
--- a/src/HOL/Import/import_rule.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Import/import_rule.ML Wed Mar 04 19:53:18 2015 +0100
@@ -44,7 +44,7 @@
fun meta_mp th1 th2 =
let
val th1a = implies_elim_all th1
- val th1b = Thm.implies_intr (strip_imp_concl (cprop_of th2)) th1a
+ val th1b = Thm.implies_intr (strip_imp_concl (Thm.cprop_of th2)) th1a
val th2a = implies_elim_all th2
val th3 = Thm.implies_elim th1b th2a
in
@@ -53,8 +53,8 @@
fun meta_eq_to_obj_eq th =
let
- val (tml, tmr) = Thm.dest_binop (strip_imp_concl (cprop_of th))
- val cty = ctyp_of_term tml
+ val (tml, tmr) = Thm.dest_binop (strip_imp_concl (Thm.cprop_of th))
+ val cty = Thm.ctyp_of_term tml
val i = Drule.instantiate' [SOME cty] [SOME tml, SOME tmr]
@{thm meta_eq_to_obj_eq}
in
@@ -65,7 +65,7 @@
fun eq_mp th1 th2 =
let
- val (tm1l, tm1r) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th1)))
+ val (tm1l, tm1r) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1)))
val i1 = Drule.instantiate' [] [SOME tm1l, SOME tm1r] @{thm iffD1}
val i2 = meta_mp i1 th1
in
@@ -74,11 +74,11 @@
fun comb th1 th2 =
let
- val t1c = Thm.dest_arg (strip_imp_concl (cprop_of th1))
- val t2c = Thm.dest_arg (strip_imp_concl (cprop_of th2))
+ val t1c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1))
+ val t2c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th2))
val (cf, cg) = Thm.dest_binop t1c
val (cx, cy) = Thm.dest_binop t2c
- val [fd, fr] = Thm.dest_ctyp (ctyp_of_term cf)
+ val [fd, fr] = Thm.dest_ctyp (Thm.ctyp_of_term cf)
val i1 = Drule.instantiate' [SOME fd, SOME fr]
[SOME cf, SOME cg, SOME cx, SOME cy] @{thm cong}
val i2 = meta_mp i1 th1
@@ -88,11 +88,11 @@
fun trans th1 th2 =
let
- val t1c = Thm.dest_arg (strip_imp_concl (cprop_of th1))
- val t2c = Thm.dest_arg (strip_imp_concl (cprop_of th2))
+ val t1c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1))
+ val t2c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th2))
val (r, s) = Thm.dest_binop t1c
val (_, t) = Thm.dest_binop t2c
- val ty = ctyp_of_term r
+ val ty = Thm.ctyp_of_term r
val i1 = Drule.instantiate' [SOME ty] [SOME r, SOME s, SOME t] @{thm trans}
val i2 = meta_mp i1 th1
in
@@ -101,17 +101,17 @@
fun deduct th1 th2 =
let
- val th1c = strip_imp_concl (cprop_of th1)
- val th2c = strip_imp_concl (cprop_of th2)
+ val th1c = strip_imp_concl (Thm.cprop_of th1)
+ val th2c = strip_imp_concl (Thm.cprop_of th2)
val th1a = implies_elim_all th1
val th2a = implies_elim_all th2
val th1b = Thm.implies_intr th2c th1a
val th2b = Thm.implies_intr th1c th2a
val i = Drule.instantiate' []
[SOME (Thm.dest_arg th1c), SOME (Thm.dest_arg th2c)] @{thm iffI}
- val i1 = Thm.implies_elim i (Thm.assume (cprop_of th2b))
+ val i1 = Thm.implies_elim i (Thm.assume (Thm.cprop_of th2b))
val i2 = Thm.implies_elim i1 th1b
- val i3 = Thm.implies_intr (cprop_of th2b) i2
+ val i3 = Thm.implies_intr (Thm.cprop_of th2b) i2
val i4 = Thm.implies_elim i3 th2b
in
implies_intr_hyps i4
@@ -119,7 +119,7 @@
fun conj1 th =
let
- val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th)))
+ val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th)))
val i = Drule.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct1}
in
meta_mp i th
@@ -127,7 +127,7 @@
fun conj2 th =
let
- val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th)))
+ val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th)))
val i = Drule.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct2}
in
meta_mp i th
@@ -143,7 +143,7 @@
fun abs cv th =
let
val th1 = implies_elim_all th
- val (tl, tr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (cprop_of th1)))
+ val (tl, tr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1)))
val (ll, lr) = (Thm.lambda cv tl, Thm.lambda cv tr)
val (al, ar) = (Thm.apply ll cv, Thm.apply lr cv)
val bl = beta al
@@ -151,7 +151,7 @@
val th2 = trans (trans bl th1) br
val th3 = implies_elim_all th2
val th4 = Thm.forall_intr cv th3
- val i = Drule.instantiate' [SOME (ctyp_of_term cv), SOME (ctyp_of_term tl)]
+ val i = Drule.instantiate' [SOME (Thm.ctyp_of_term cv), SOME (Thm.ctyp_of_term tl)]
[SOME ll, SOME lr] @{thm ext2}
in
meta_mp i th4
@@ -159,18 +159,18 @@
fun freezeT thm =
let
- val tvars = Term.add_tvars (prop_of thm) []
+ val tvars = Term.add_tvars (Thm.prop_of thm) []
val tfrees = map (fn ((t, _), s) => TFree (t, s)) tvars
val tvars = map TVar tvars
val thy = Thm.theory_of_thm thm
- fun inst ty = ctyp_of thy ty
+ fun inst ty = Thm.ctyp_of thy ty
in
Thm.instantiate ((map inst tvars ~~ map inst tfrees), []) thm
end
fun def' constname rhs thy =
let
- val rhs = term_of rhs
+ val rhs = Thm.term_of rhs
val typ = type_of rhs
val constbinding = Binding.name constname
val thy1 = Sign.add_consts [(constbinding, typ, NoSyn)] thy
@@ -199,23 +199,24 @@
fun typedef_hollight th thy =
let
- val (th_s, cn) = Thm.dest_comb (Thm.dest_arg (cprop_of th))
+ val (th_s, cn) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th))
val (th_s, abst) = Thm.dest_comb th_s
val rept = Thm.dest_arg th_s
val P = Thm.dest_arg cn
- val [nty, oty] = Thm.dest_ctyp (ctyp_of_term rept)
+ val [nty, oty] = Thm.dest_ctyp (Thm.ctyp_of_term rept)
in
Drule.instantiate' [SOME nty, SOME oty] [SOME rept, SOME abst, SOME P,
- SOME (cterm_of thy (Free ("a", typ_of nty))),
- SOME (cterm_of thy (Free ("r", typ_of oty)))] @{thm typedef_hol2hollight}
+ SOME (Thm.cterm_of thy (Free ("a", Thm.typ_of nty))),
+ SOME (Thm.cterm_of thy (Free ("r", Thm.typ_of oty)))] @{thm typedef_hol2hollight}
end
fun tydef' tycname abs_name rep_name cP ct td_th thy =
let
- val ctT = ctyp_of_term ct
+ val ctT = Thm.ctyp_of_term ct
val nonempty = Drule.instantiate' [SOME ctT] [SOME cP, SOME ct] @{thm light_ex_imp_nonempty}
val th2 = meta_mp nonempty td_th
- val c = case concl_of th2 of
+ val c =
+ case Thm.concl_of th2 of
_ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
| _ => error "type_introduction: bad type definition theorem"
val tfrees = Term.add_tfrees c []
@@ -225,15 +226,15 @@
(SOME (Binding.name rep_name, Binding.name abs_name)) (fn _ => rtac th2 1) thy
val aty = #abs_type (#1 typedef_info)
val th = freezeT (#type_definition (#2 typedef_info))
- val (th_s, _) = Thm.dest_comb (Thm.dest_arg (cprop_of th))
+ val (th_s, _) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th))
val (th_s, abst) = Thm.dest_comb th_s
val rept = Thm.dest_arg th_s
- val [nty, oty] = Thm.dest_ctyp (ctyp_of_term rept)
+ val [nty, oty] = Thm.dest_ctyp (Thm.ctyp_of_term rept)
val typedef_th =
Drule.instantiate'
[SOME nty, SOME oty]
- [SOME rept, SOME abst, SOME cP, SOME (cterm_of thy' (Free ("a", aty))),
- SOME (cterm_of thy' (Free ("r", typ_of ctT)))]
+ [SOME rept, SOME abst, SOME cP, SOME (Thm.cterm_of thy' (Free ("a", aty))),
+ SOME (Thm.cterm_of thy' (Free ("r", Thm.typ_of ctT)))]
@{thm typedef_hol2hollight}
val th4 = typedef_th OF [#type_definition (#2 typedef_info)]
in
@@ -259,14 +260,14 @@
let
fun assoc _ [] = error "assoc"
| assoc x ((x',y)::rest) = if x = x' then y else assoc x rest
- val lambda = map (fn (a, b) => (typ_of a, b)) lambda
- val tys_before = Term.add_tfrees (prop_of th) []
+ val lambda = map (fn (a, b) => (Thm.typ_of a, b)) lambda
+ val tys_before = Term.add_tfrees (Thm.prop_of th) []
val th1 = Thm.varifyT_global th
- val tys_after = Term.add_tvars (prop_of th1) []
+ val tys_after = Term.add_tvars (Thm.prop_of th1) []
val tyinst = map2 (fn bef => fn iS =>
(case try (assoc (TFree bef)) lambda of
- SOME cty => (ctyp_of thy (TVar iS), cty)
- | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
+ SOME cty => (Thm.ctyp_of thy (TVar iS), cty)
+ | NONE => (Thm.ctyp_of thy (TVar iS), Thm.ctyp_of thy (TFree bef))
)) tys_before tys_after
in
Thm.instantiate (tyinst,[]) th1
@@ -328,12 +329,12 @@
fun store_thm binding thm thy =
let
val thm = Drule.export_without_context_open thm
- val tvs = Term.add_tvars (prop_of thm) []
+ val tvs = Term.add_tvars (Thm.prop_of thm) []
val tns = map (fn (_, _) => "'") tvs
val nms = fst (fold_map Name.variant tns (Variable.names_of (Proof_Context.init_global thy)))
val vs = map TVar ((nms ~~ (map (snd o fst) tvs)) ~~ (map snd tvs))
- val cvs = map (ctyp_of thy) vs
- val ctvs = map (ctyp_of thy) (map TVar tvs)
+ val cvs = map (Thm.ctyp_of thy) vs
+ val ctvs = map (Thm.ctyp_of thy) (map TVar tvs)
val thm' = Thm.instantiate ((ctvs ~~ cvs), []) thm
in
snd (Global_Theory.add_thm ((binding, thm'), []) thy)
@@ -410,14 +411,14 @@
end
| process (thy, state) (#"Y", [name, _, _]) = setth (mtydef name thy) (thy, state)
| process (thy, state) (#"t", [n]) =
- setty (ctyp_of thy (TFree ("'" ^ (transl_qm n), @{sort type}))) (thy, state)
+ setty (Thm.ctyp_of thy (TFree ("'" ^ (transl_qm n), @{sort type}))) (thy, state)
| process (thy, state) (#"a", n :: l) =
fold_map getty l (thy, state) |>>
- (fn tys => ctyp_of thy (Type (gettyname n thy, map typ_of tys))) |-> setty
+ (fn tys => Thm.ctyp_of thy (Type (gettyname n thy, map Thm.typ_of tys))) |-> setty
| process (thy, state) (#"v", [n, ty]) =
- getty ty (thy, state) |>> (fn ty => cterm_of thy (Free (transl_dot n, typ_of ty))) |-> settm
+ getty ty (thy, state) |>> (fn ty => Thm.cterm_of thy (Free (transl_dot n, Thm.typ_of ty))) |-> settm
| process (thy, state) (#"c", [n, ty]) =
- getty ty (thy, state) |>> (fn ty => cterm_of thy (Const (getconstname n thy, typ_of ty))) |-> settm
+ getty ty (thy, state) |>> (fn ty => Thm.cterm_of thy (Const (getconstname n thy, Thm.typ_of ty))) |-> settm
| process tstate (#"f", [t1, t2]) =
gettm t1 tstate ||>> gettm t2 |>> (fn (t1, t2) => Thm.apply t1 t2) |-> settm
| process tstate (#"l", [t1, t2]) =
--- a/src/HOL/Int.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Int.thy Wed Mar 04 19:53:18 2015 +0100
@@ -753,7 +753,7 @@
simproc_setup fast_arith ("(m::'a::linordered_idom) < n" |
"(m::'a::linordered_idom) <= n" |
"(m::'a::linordered_idom) = n") =
- {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (term_of ct) *}
+ {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
subsection{*More Inequality Reasoning*}
--- a/src/HOL/Library/Code_Abstract_Nat.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Code_Abstract_Nat.thy Wed Mar 04 19:53:18 2015 +0100
@@ -59,10 +59,10 @@
val thy = Proof_Context.theory_of ctxt;
val vname = singleton (Name.variant_list (map fst
(fold (Term.add_var_names o Thm.full_prop_of) thms []))) "n";
- val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
- val lhs_of = snd o Thm.dest_comb o fst o Thm.dest_comb o cprop_of;
- val rhs_of = snd o Thm.dest_comb o cprop_of;
- fun find_vars ct = (case term_of ct of
+ val cv = Thm.cterm_of thy (Var ((vname, 0), HOLogic.natT));
+ val lhs_of = snd o Thm.dest_comb o fst o Thm.dest_comb o Thm.cprop_of;
+ val rhs_of = snd o Thm.dest_comb o Thm.cprop_of;
+ fun find_vars ct = (case Thm.term_of ct of
(Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
| _ $ _ =>
let val (ct1, ct2) = Thm.dest_comb ct
@@ -79,7 +79,7 @@
Thm.implies_elim
(Conv.fconv_rule (Thm.beta_conversion true)
(Drule.instantiate'
- [SOME (ctyp_of_term ct)] [SOME (Thm.lambda cv ct),
+ [SOME (Thm.ctyp_of_term ct)] [SOME (Thm.lambda cv ct),
SOME (Thm.lambda cv' (rhs_of thm)), NONE, SOME cv']
Suc_if_eq)) (Thm.forall_intr cv' thm)
in
@@ -97,7 +97,7 @@
fun eqn_suc_base_preproc ctxt thms =
let
- val dest = fst o Logic.dest_equals o prop_of;
+ val dest = fst o Logic.dest_equals o Thm.prop_of;
val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
in
if forall (can dest) thms andalso exists (contains_suc o dest) thms
--- a/src/HOL/Library/Countable.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Countable.thy Wed Mar 04 19:53:18 2015 +0100
@@ -171,7 +171,7 @@
val typedef_info = hd (Typedef.get_info ctxt ty_name)
val typedef_thm = #type_definition (snd typedef_info)
val pred_name =
- (case HOLogic.dest_Trueprop (concl_of typedef_thm) of
+ (case HOLogic.dest_Trueprop (Thm.concl_of typedef_thm) of
(_ $ _ $ _ $ (_ $ Const (n, _))) => n
| _ => raise Match)
val induct_info = Inductive.the_inductive ctxt pred_name
--- a/src/HOL/Library/Extended_Nat.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Extended_Nat.thy Wed Mar 04 19:53:18 2015 +0100
@@ -561,15 +561,15 @@
simproc_setup enat_eq_cancel
("(l::enat) + m = n" | "(l::enat) = m + n") =
- {* fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (term_of ct) *}
+ {* fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
simproc_setup enat_le_cancel
("(l::enat) + m \<le> n" | "(l::enat) \<le> m + n") =
- {* fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (term_of ct) *}
+ {* fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
simproc_setup enat_less_cancel
("(l::enat) + m < n" | "(l::enat) < m + n") =
- {* fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (term_of ct) *}
+ {* fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct) *}
text {* TODO: add regression tests for these simprocs *}
--- a/src/HOL/Library/Reflection.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Reflection.thy Wed Mar 04 19:53:18 2015 +0100
@@ -23,7 +23,7 @@
val rulesN = "rules";
val any_keyword = keyword onlyN || keyword rulesN;
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
- val terms = thms >> map (term_of o Drule.dest_term);
+ val terms = thms >> map (Thm.term_of o Drule.dest_term);
in
thms -- Scan.optional (keyword rulesN |-- thms) [] --
Scan.option (keyword onlyN |-- Args.term) >>
--- a/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML Wed Mar 04 19:53:18 2015 +0100
@@ -31,7 +31,7 @@
fun string_of_varpow x k =
let
- val term = term_of x
+ val term = Thm.term_of x
val name =
(case term of
Free (n, _) => n
--- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Wed Mar 04 19:53:18 2015 +0100
@@ -234,7 +234,7 @@
val div_tm = @{cterm "op / :: real => _"}
val pow_tm = @{cterm "op ^ :: real => _"}
val zero_tm = @{cterm "0:: real"}
- val is_numeral = can (HOLogic.dest_number o term_of)
+ val is_numeral = can (HOLogic.dest_number o Thm.term_of)
fun poly_of_term tm =
if tm aconvc zero_tm then poly_0
else
@@ -256,7 +256,7 @@
val (opr,l) = Thm.dest_comb lop
in
if opr aconvc pow_tm andalso is_numeral r
- then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o term_of) r)
+ then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o Thm.term_of) r)
else if opr aconvc add_tm
then poly_add (poly_of_term l) (poly_of_term r)
else if opr aconvc sub_tm
@@ -278,7 +278,7 @@
end handle CTERM ("dest_comb",_) => poly_var tm)
in
val poly_of_term = fn tm =>
- if type_of (term_of tm) = @{typ real}
+ if type_of (Thm.term_of tm) = @{typ real}
then poly_of_term tm
else error "poly_of_term: term does not have real type"
end;
@@ -779,8 +779,8 @@
(* Interface to HOL. *)
local
open Conv
- val concl = Thm.dest_arg o cprop_of
- fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
+ val concl = Thm.dest_arg o Thm.cprop_of
+ fun simple_cterm_ord t u = Term_Ord.fast_term_ord (Thm.term_of t, Thm.term_of u) = LESS
in
(* FIXME: Replace tryfind by get_first !! *)
fun real_nonlinear_prover proof_method ctxt =
@@ -881,8 +881,8 @@
local
open Conv
- fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
- val concl = Thm.dest_arg o cprop_of
+ fun simple_cterm_ord t u = Term_Ord.fast_term_ord (Thm.term_of t, Thm.term_of u) = LESS
+ val concl = Thm.dest_arg o Thm.cprop_of
val shuffle1 =
fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))"
by (atomize (full)) (simp add: field_simps)})
@@ -890,7 +890,7 @@
fconv_rule (rewr_conv @{lemma "(x + a == y) == (x == y - (a::real))"
by (atomize (full)) (simp add: field_simps)})
fun substitutable_monomial fvs tm =
- (case term_of tm of
+ (case Thm.term_of tm of
Free (_, @{typ real}) =>
if not (member (op aconvc) fvs tm) then (Rat.one, tm)
else raise Failure "substitutable_monomial"
@@ -907,11 +907,11 @@
fun isolate_variable v th =
let
- val w = Thm.dest_arg1 (cprop_of th)
+ val w = Thm.dest_arg1 (Thm.cprop_of th)
in
if v aconvc w then th
else
- (case term_of w of
+ (case Thm.term_of w of
@{term "op + :: real => _"} $ _ $ _ =>
if Thm.dest_arg1 w aconvc v then shuffle2 th
else isolate_variable v (shuffle1 th)
@@ -951,7 +951,7 @@
in
substfirst
(filter_out
- (fn t => (Thm.dest_arg1 o Thm.dest_arg o cprop_of) t aconvc @{cterm "0::real"})
+ (fn t => (Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of) t aconvc @{cterm "0::real"})
(map modify eqs),
map modify les,
map modify lts)
@@ -986,7 +986,7 @@
fun check_sos kcts ct =
let
- val t = term_of ct
+ val t = Thm.term_of ct
val _ =
if not (null (Term.add_tfrees t []) andalso null (Term.add_tvars t []))
then error "SOS: not sos. Additional type varables"
@@ -1022,10 +1022,10 @@
local
- val is_numeral = can (HOLogic.dest_number o term_of)
+ val is_numeral = can (HOLogic.dest_number o Thm.term_of)
in
fun get_denom b ct =
- (case term_of ct of
+ (case Thm.term_of ct of
@{term "op / :: real => _"} $ _ $ _ =>
if is_numeral (Thm.dest_arg ct)
then get_denom b (Thm.dest_arg1 ct)
--- a/src/HOL/Library/positivstellensatz.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/positivstellensatz.ML Wed Mar 04 19:53:18 2015 +0100
@@ -64,7 +64,7 @@
structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
structure Termfunc = FuncFun(type key = term val ord = Term_Ord.fast_term_ord);
-val cterm_ord = Term_Ord.fast_term_ord o apply2 term_of
+val cterm_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of
structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);
@@ -170,11 +170,11 @@
(* Some useful derived rules *)
fun deduct_antisym_rule tha thb =
- Thm.equal_intr (Thm.implies_intr (cprop_of thb) tha)
- (Thm.implies_intr (cprop_of tha) thb);
+ Thm.equal_intr (Thm.implies_intr (Thm.cprop_of thb) tha)
+ (Thm.implies_intr (Thm.cprop_of tha) thb);
fun prove_hyp tha thb =
- if exists (curry op aconv (concl_of tha)) (Thm.hyps_of thb) (* FIXME !? *)
+ if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb) (* FIXME !? *)
then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;
val pth = @{lemma "(((x::real) < y) == (y - x > 0))" and "((x <= y) == (y - x >= 0))" and
@@ -279,10 +279,10 @@
fun literals_conv bops uops cv =
let
fun h t =
- case (term_of t) of
+ (case Thm.term_of t of
b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t
| u$_ => if member (op aconv) uops u then arg_conv h t else cv t
- | _ => cv t
+ | _ => cv t)
in h end;
fun cterm_of_rat x =
@@ -296,9 +296,9 @@
end;
fun dest_ratconst t =
- case term_of t of
+ case Thm.term_of t of
Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
- | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
+ | _ => Rat.rat_of_int (HOLogic.dest_number (Thm.term_of t) |> snd)
fun is_ratconst t = can dest_ratconst t
(*
@@ -311,14 +311,14 @@
fun find_cterm p t =
if p t then t else
- case term_of t of
+ case Thm.term_of t of
_$_ => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t))
| Abs (_,_,_) => find_cterm p (Thm.dest_abs NONE t |> snd)
| _ => raise CTERM ("find_cterm",[t]);
(* Some conversions-related stuff which has been forbidden entrance into Pure/conv.ML*)
fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
-fun is_comb t = case (term_of t) of _$_ => true | _ => false;
+fun is_comb t = (case Thm.term_of t of _ $ _ => true | _ => false);
fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))
handle CTERM _ => false;
@@ -428,7 +428,7 @@
nnf_conv ctxt then_conv skolemize_conv then_conv prenex_conv then_conv
weak_dnf_conv
- val concl = Thm.dest_arg o cprop_of
+ val concl = Thm.dest_arg o Thm.cprop_of
fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false)
val is_req = is_binop @{cterm "op =:: real => _"}
val is_ge = is_binop @{cterm "op <=:: real => _"}
@@ -501,36 +501,36 @@
val strip_exists =
let
fun h (acc, t) =
- case (term_of t) of
+ case Thm.term_of t of
Const(@{const_name Ex},_)$Abs(_,_,_) =>
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
| _ => (acc,t)
in fn t => h ([],t)
end
fun name_of x =
- case term_of x of
+ case Thm.term_of x of
Free(s,_) => s
| Var ((s,_),_) => s
| _ => "x"
fun mk_forall x th =
Drule.arg_cong_rule
- (instantiate_cterm' [SOME (ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" })
+ (instantiate_cterm' [SOME (Thm.ctyp_of_term x)] [] @{cpat "All :: (?'a => bool) => _" })
(Thm.abstract_rule (name_of x) x th)
val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
fun ext T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat Ex}
- fun mk_ex v t = Thm.apply (ext (ctyp_of_term v)) (Thm.lambda v t)
+ fun mk_ex v t = Thm.apply (ext (Thm.ctyp_of_term v)) (Thm.lambda v t)
fun choose v th th' =
- case concl_of th of
+ case Thm.concl_of th of
@{term Trueprop} $ (Const(@{const_name Ex},_)$_) =>
let
- val p = (funpow 2 Thm.dest_arg o cprop_of) th
- val T = (hd o Thm.dest_ctyp o ctyp_of_term) p
+ val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th
+ val T = (hd o Thm.dest_ctyp o Thm.ctyp_of_term) p
val th0 = fconv_rule (Thm.beta_conversion true)
- (instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o cprop_of) th'] exE)
+ (instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE)
val pv = (Thm.rhs_of o Thm.beta_conversion true)
(Thm.apply @{cterm Trueprop} (Thm.apply p v))
val th1 = Thm.forall_intr v (Thm.implies_intr pv th')
@@ -545,7 +545,7 @@
val strip_forall =
let
fun h (acc, t) =
- case (term_of t) of
+ case Thm.term_of t of
Const(@{const_name All},_)$Abs(_,_,_) =>
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
| _ => (acc,t)
@@ -693,15 +693,15 @@
end
fun is_alien ct =
- case term_of ct of
- Const(@{const_name "real"}, _)$ n =>
- if can HOLogic.dest_number n then false else true
- | _ => false
+ case Thm.term_of ct of
+ Const(@{const_name "real"}, _)$ n =>
+ if can HOLogic.dest_number n then false else true
+ | _ => false
in
fun real_linear_prover translator (eq,le,lt) =
let
- val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o cprop_of
- val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o cprop_of
+ val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of
+ val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o Thm.cprop_of
val eq_pols = map lhs eq
val le_pols = map rhs le
val lt_pols = map rhs lt
@@ -770,7 +770,7 @@
fun gen_prover_real_arith ctxt prover =
let
- fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS
+ fun simple_cterm_ord t u = Term_Ord.term_ord (Thm.term_of t, Thm.term_of u) = LESS
val {add, mul, neg, pow = _, sub = _, main} =
Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt
(the (Semiring_Normalizer.match ctxt @{cterm "(0::real) + 1"}))
--- a/src/HOL/Library/refute.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/refute.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1202,13 +1202,13 @@
fun refute_goal ctxt params th i =
let
- val t = th |> prop_of
+ val t = th |> Thm.prop_of
in
if Logic.count_prems t = 0 then
(writeln "No subgoal!"; "none")
else
let
- val assms = map term_of (Assumption.all_assms_of ctxt)
+ val assms = map Thm.term_of (Assumption.all_assms_of ctxt)
val (t, frees) = Logic.goal_params t i
in
refute_term ctxt params assms (subst_bounds (frees, t))
--- a/src/HOL/Library/simps_case_conv.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Library/simps_case_conv.ML Wed Mar 04 19:53:18 2015 +0100
@@ -27,7 +27,7 @@
#> map_filter (Ctr_Sugar.ctr_sugar_of ctxt)
#> map #split
-val strip_eq = prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq
+val strip_eq = Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq
local
@@ -161,7 +161,7 @@
fun apply_split ctxt split thm = Seq.of_list
let val ((_,thm'), ctxt') = Variable.import false [thm] ctxt in
- (Variable.export ctxt' ctxt) (filter (was_split o prop_of) (thm' RL [split]))
+ (Variable.export ctxt' ctxt) (filter (was_split o Thm.prop_of) (thm' RL [split]))
end
fun forward_tac rules t = Seq.of_list ([t] RL rules)
@@ -176,7 +176,7 @@
fun do_split ctxt split =
let
val split' = split RS iffD1;
- val split_rhs = concl_of (hd (snd (fst (Variable.import false [split'] ctxt))))
+ val split_rhs = Thm.concl_of (hd (snd (fst (Variable.import false [split'] ctxt))))
in if was_split split_rhs
then DETERM (apply_split ctxt split') THEN get_rules_once_split
else raise TERM ("malformed split rule", [split_rhs])
--- a/src/HOL/List.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/List.thy Wed Mar 04 19:53:18 2015 +0100
@@ -708,7 +708,7 @@
val inner_t =
fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t)
(rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
- val lhs = term_of redex
+ val lhs = Thm.term_of redex
val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
in
@@ -717,7 +717,7 @@
(fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
end))
in
- make_inner_eqs [] [] [] (dest_set (term_of redex))
+ make_inner_eqs [] [] [] (dest_set (Thm.term_of redex))
end
end
@@ -1043,7 +1043,7 @@
else if lastl aconv lastr then rearr @{thm append_same_eq}
else NONE
end;
- in fn _ => fn ctxt => fn ct => list_eq ctxt (term_of ct) end;
+ in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end;
*}
--- a/src/HOL/Matrix_LP/ComputeHOL.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Matrix_LP/ComputeHOL.thy Wed Mar 04 19:53:18 2015 +0100
@@ -148,7 +148,7 @@
struct
local
-fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq));
+fun lhs_of eq = fst (Thm.dest_equals (Thm.cprop_of eq));
in
fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [ct])
| rewrite_conv (eq :: eqs) ct =
--- a/src/HOL/Matrix_LP/Compute_Oracle/compute.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML Wed Mar 04 19:53:18 2015 +0100
@@ -368,7 +368,7 @@
fun export_thm thy hyps shyps prop =
let
val th = export_oracle (thy, hyps, shyps, prop)
- val hyps = map (fn h => Thm.assume (cterm_of thy h)) hyps
+ val hyps = map (fn h => Thm.assume (Thm.cterm_of thy h)) hyps
in
fold (fn h => fn p => Thm.implies_elim p h) hyps th
end
@@ -378,7 +378,7 @@
fun rewrite computer ct =
let
val thy = Thm.theory_of_cterm ct
- val {t=t',T=ty,...} = rep_cterm ct
+ val {t=t',T=ty,...} = Thm.rep_cterm ct
val _ = super_theory (theory_of computer) thy
val naming = naming_of computer
val (encoding, t) = remove_types (encoding_of computer) t'
@@ -401,7 +401,7 @@
fun make_theorem computer th vars =
let
- val _ = super_theory (theory_of computer) (theory_of_thm th)
+ val _ = super_theory (theory_of computer) (Thm.theory_of_thm th)
val (ComputeThm (hyps, shyps, prop)) = thm2cthm th
@@ -452,7 +452,7 @@
val (encoding, concl) = remove_types encoding (Logic.strip_imp_concl prop)
val _ = set_encoding computer encoding
in
- Theorem (theory_of_thm th, stamp_of computer, vartab, varsubst,
+ Theorem (Thm.theory_of_thm th, stamp_of computer, vartab, varsubst,
prems, concl, hyps, shyps)
end
@@ -506,8 +506,8 @@
fun compute_inst (s, ct) vs =
let
- val _ = super_theory (theory_of_cterm ct) thy
- val ty = typ_of (ctyp_of_term ct)
+ val _ = super_theory (Thm.theory_of_cterm ct) thy
+ val ty = Thm.typ_of (Thm.ctyp_of_term ct)
in
(case Symtab.lookup vartab s of
NONE => raise Compute ("instantiate: variable '"^s^"' not found in theorem")
@@ -519,7 +519,7 @@
raise Compute ("instantiate: wrong type for variable '"^s^"'")
else
let
- val t = rewrite computer (term_of ct)
+ val t = rewrite computer (Thm.term_of ct)
val _ = assert_varfree vs t
val _ = assert_closed t
in
@@ -611,7 +611,7 @@
val thy =
let
val thy1 = theory_of_theorem th
- val thy2 = theory_of_thm th'
+ val thy2 = Thm.theory_of_thm th'
in
if Theory.subthy (thy1, thy2) then thy2
else if Theory.subthy (thy2, thy1) then thy1 else
--- a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML Wed Mar 04 19:53:18 2015 +0100
@@ -246,7 +246,7 @@
fun collect_consts_of_thm th =
let
- val th = prop_of th
+ val th = Thm.prop_of th
val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th)
val (left, right) = Logic.dest_equals th
in
@@ -315,7 +315,8 @@
end
fun conv_subst thy (subst : Type.tyenv) =
- map (fn (iname, (sort, ty)) => (ctyp_of thy (TVar (iname, sort)), ctyp_of thy ty)) (Vartab.dest subst)
+ map (fn (iname, (sort, ty)) => (Thm.ctyp_of thy (TVar (iname, sort)), Thm.ctyp_of thy ty))
+ (Vartab.dest subst)
fun add_monos thy monocs pats ths =
let
@@ -433,7 +434,7 @@
fun rewrite pc cts =
let
- val _ = add_instances' pc (map term_of cts)
+ val _ = add_instances' pc (map Thm.term_of cts)
val computer = (computer_of pc)
in
map (fn ct => Compute.rewrite computer ct) cts
@@ -443,7 +444,7 @@
fun make_theorem pc th vars =
let
- val _ = add_instances' pc [prop_of th]
+ val _ = add_instances' pc [Thm.prop_of th]
in
Compute.make_theorem (computer_of pc) th vars
@@ -451,7 +452,7 @@
fun instantiate pc insts th =
let
- val _ = add_instances' pc (map (term_of o snd) insts)
+ val _ = add_instances' pc (map (Thm.term_of o snd) insts)
in
Compute.instantiate (computer_of pc) insts th
end
@@ -460,7 +461,7 @@
fun modus_ponens pc prem_no th' th =
let
- val _ = add_instances' pc [prop_of th']
+ val _ = add_instances' pc [Thm.prop_of th']
in
Compute.modus_ponens (computer_of pc) prem_no th' th
end
--- a/src/HOL/Matrix_LP/matrixlp.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Matrix_LP/matrixlp.ML Wed Mar 04 19:53:18 2015 +0100
@@ -23,7 +23,7 @@
fun matrix_simplify th =
let
- val simp_th = matrix_compute (cprop_of th)
+ val simp_th = matrix_compute (Thm.cprop_of th)
val th = Thm.strip_shyps (Thm.equal_elim simp_th th)
fun removeTrue th = removeTrue (Thm.implies_elim th TrueI) handle THM _ => th
in
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -310,7 +310,7 @@
fun get_prover ctxt name params goal all_facts =
let
- val learn = Sledgehammer_MaSh.mash_learn_proof ctxt params (prop_of goal) all_facts
+ val learn = Sledgehammer_MaSh.mash_learn_proof ctxt params (Thm.prop_of goal) all_facts
in
Sledgehammer_Prover_Minimize.get_minimizing_prover ctxt Sledgehammer_Prover.Normal learn name
end
--- a/src/HOL/Multivariate_Analysis/normarith.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Multivariate_Analysis/normarith.ML Wed Mar 04 19:53:18 2015 +0100
@@ -15,15 +15,15 @@
open Conv;
val bool_eq = op = : bool *bool -> bool
- fun dest_ratconst t = case term_of t of
+ fun dest_ratconst t = case Thm.term_of t of
Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
| Const(@{const_name inverse}, _)$a => Rat.rat_of_quotient(1, HOLogic.dest_number a |> snd)
- | _ => Rat.rat_of_int (HOLogic.dest_number (term_of t) |> snd)
+ | _ => Rat.rat_of_int (HOLogic.dest_number (Thm.term_of t) |> snd)
fun is_ratconst t = can dest_ratconst t
- fun augment_norm b t acc = case term_of t of
+ fun augment_norm b t acc = case Thm.term_of t of
Const(@{const_name norm}, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc
| _ => acc
- fun find_normedterms t acc = case term_of t of
+ fun find_normedterms t acc = case Thm.term_of t of
@{term "op + :: real => _"}$_$_ =>
find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc)
| @{term "op * :: real => _"}$_$_ =>
@@ -50,7 +50,7 @@
fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
*)
-fun vector_lincomb t = case term_of t of
+fun vector_lincomb t = case Thm.term_of t of
Const(@{const_name plus}, _) $ _ $ _ =>
cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t))
| Const(@{const_name minus}, _) $ _ $ _ =>
@@ -81,7 +81,7 @@
end
| SOME _ => fns) ts []
-fun replacenegnorms cv t = case term_of t of
+fun replacenegnorms cv t = case Thm.term_of t of
@{term "op + :: real => _"}$_$_ => binop_conv (replacenegnorms cv) t
| @{term "op * :: real => _"}$_$_ =>
if dest_ratconst (Thm.dest_arg1 t) </ Rat.zero then arg_conv cv t else Thm.reflexive t
@@ -203,7 +203,7 @@
handle CTERM _ =>
(apply_pth8 ctxt ct
handle CTERM _ =>
- (case term_of ct of
+ (case Thm.term_of ct of
Const(@{const_name plus},_)$lt$rt =>
let
val l = headvector lt
@@ -219,7 +219,7 @@
end
| _ => Thm.reflexive ct))
-fun vector_canon_conv ctxt ct = case term_of ct of
+fun vector_canon_conv ctxt ct = case Thm.term_of ct of
Const(@{const_name plus},_)$_$_ =>
let
val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb
@@ -248,7 +248,7 @@
*)
| _ => apply_pth1 ct
-fun norm_canon_conv ctxt ct = case term_of ct of
+fun norm_canon_conv ctxt ct = case Thm.term_of ct of
Const(@{const_name norm},_)$_ => arg_conv (vector_canon_conv ctxt) ct
| _ => raise CTERM ("norm_canon_conv", [ct])
@@ -258,9 +258,9 @@
local
val pth_zero = @{thm norm_zero}
- val tv_n = (ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o cprop_of)
+ val tv_n = (Thm.ctyp_of_term o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of)
pth_zero
- val concl = Thm.dest_arg o cprop_of
+ val concl = Thm.dest_arg o Thm.cprop_of
fun real_vector_combo_prover ctxt translator (nubs,ges,gts) =
let
(* FIXME: Should be computed statically!!*)
@@ -319,7 +319,7 @@
(fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
in fst (RealArith.real_linear_prover translator
- (map (fn t => Drule.instantiate_normalize ([(tv_n, ctyp_of_term t)],[]) pth_zero)
+ (map (fn t => Drule.instantiate_normalize ([(tv_n, Thm.ctyp_of_term t)],[]) pth_zero)
zerodests,
map (fconv_rule (try_conv (Conv.top_sweep_conv (K (norm_canon_conv ctxt)) ctxt) then_conv
arg_conv (arg_conv real_poly_conv))) ges',
@@ -332,7 +332,7 @@
local
val pth = @{thm norm_imp_pos_and_ge}
val norm_mp = match_mp pth
- val concl = Thm.dest_arg o cprop_of
+ val concl = Thm.dest_arg o Thm.cprop_of
fun conjunct1 th = th RS @{thm conjunct1}
fun conjunct2 th = th RS @{thm conjunct2}
fun real_vector_ineq_prover ctxt translator (ges,gts) =
@@ -342,8 +342,8 @@
val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
- fun mk_norm t = Thm.apply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
- fun mk_equals l r = Thm.apply (Thm.apply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
+ fun mk_norm t = Thm.apply (instantiate_cterm' [SOME (Thm.ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
+ fun mk_equals l r = Thm.apply (Thm.apply (instantiate_cterm' [SOME (Thm.ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
val asl = map2 (fn (t,_) => fn n => Thm.assume (mk_equals (mk_norm t) (Proof_Context.cterm_of ctxt' (Free(n,@{typ real}))))) lctab fxns
val replace_conv = try_conv (rewrs_conv asl)
val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv))
@@ -353,7 +353,7 @@
val gts' = map replace_rule gts
val nubs = map (conjunct2 o norm_mp) asl
val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts')
- val shs = filter (member (fn (t,th) => t aconvc cprop_of th) asl) (#hyps (crep_thm th1))
+ val shs = filter (member (fn (t,th) => t aconvc Thm.cprop_of th) asl) (#hyps (Thm.crep_thm th1))
val th11 = hd (Variable.export ctxt' ctxt [fold Thm.implies_intr shs th1])
val cps = map (swap o Thm.dest_equals) (cprems_of th11)
val th12 = Drule.instantiate_normalize ([], cps) th11
@@ -366,7 +366,7 @@
local
val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0}))
fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
- fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
+ fun simple_cterm_ord t u = Term_Ord.term_ord (Thm.term_of t, Thm.term_of u) = LESS;
(* FIXME: Lookup in the context every time!!! Fix this !!!*)
fun splitequation ctxt th acc =
let
@@ -391,7 +391,7 @@
fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
fun norm_arith ctxt ct =
let
- val ctxt' = Variable.declare_term (term_of ct) ctxt
+ val ctxt' = Variable.declare_term (Thm.term_of ct) ctxt
val th = init_conv ctxt' ct
in Thm.equal_elim (Drule.arg_cong_rule @{cterm Trueprop} (Thm.symmetric th))
(pure ctxt' (Thm.rhs_of th))
--- a/src/HOL/Mutabelle/mutabelle_extra.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML Wed Mar 04 19:53:18 2015 +0100
@@ -236,7 +236,7 @@
val forbidden_consts = [@{const_name Pure.type}]
fun is_forbidden_theorem (s, th) =
- let val consts = Term.add_const_names (prop_of th) [] in
+ let val consts = Term.add_const_names (Thm.prop_of th) [] in
exists (member (op =) (Long_Name.explode s)) forbidden_thms orelse
exists (member (op =) forbidden_consts) consts orelse
length (Long_Name.explode s) <> 2 orelse
@@ -325,7 +325,7 @@
(fst (Variable.import_terms true [t] ctxt)))
end
-fun is_executable_thm thy th = is_executable_term thy (prop_of th)
+fun is_executable_thm thy th = is_executable_term thy (Thm.prop_of th)
val freezeT =
map_types (map_type_tvar (fn ((a, i), S) =>
@@ -333,7 +333,7 @@
fun thms_of all thy =
filter
- (fn th => (all orelse Context.theory_name (theory_of_thm th) = Context.theory_name thy)
+ (fn th => (all orelse Context.theory_name (Thm.theory_of_thm th) = Context.theory_name thy)
(* andalso is_executable_thm thy th *))
(map snd (filter_out is_forbidden_theorem (Global_Theory.all_thms_of thy false)))
@@ -368,7 +368,7 @@
map (fn (mtd_name, invoke_mtd) =>
(mtd_name, safe_invoke_mtd thy (mtd_name, invoke_mtd) mutant)) mtds)
in
- (Thm.get_name_hint thm, exec, prop_of thm, map create_mutant_subentry mutants)
+ (Thm.get_name_hint thm, exec, Thm.prop_of thm, map create_mutant_subentry mutants)
end
(* (theory -> thm -> bool -> term list -> mtd list -> 'a) -> theory -> mtd list -> thm -> 'a *)
--- a/src/HOL/NSA/HyperDef.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/NSA/HyperDef.thy Wed Mar 04 19:53:18 2015 +0100
@@ -338,7 +338,7 @@
*}
simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) <= n" | "(m::hypreal) = n") =
- {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (term_of ct) *}
+ {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
subsection {* Exponentials on the Hyperreals *}
--- a/src/HOL/NSA/NSA.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/NSA/NSA.thy Wed Mar 04 19:53:18 2015 +0100
@@ -656,7 +656,8 @@
("0 @= x" | "1 @= y" | "numeral w @= z" | "- 1 @= y" | "- numeral w @= r") =
{*
let val rule = @{thm approx_reorient} RS eq_reflection
- fun proc phi ss ct = case term_of ct of
+ fun proc phi ss ct =
+ case Thm.term_of ct of
_ $ t $ u => if can HOLogic.dest_number u then NONE
else if can HOLogic.dest_number t then SOME rule else NONE
| _ => NONE
--- a/src/HOL/NSA/transfer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/NSA/transfer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -57,7 +57,7 @@
val meta = Local_Defs.meta_rewrite_rule ctxt;
val ths' = map meta ths;
val unfolds' = map meta unfolds and refolds' = map meta refolds;
- val (_$_$t') = concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (cterm_of thy t))
+ val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of thy t))
val u = unstar_term consts t'
val tac =
rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN
@@ -73,7 +73,7 @@
(fn th =>
let
val tr = transfer_thm_of ctxt ths t
- val (_$l$r) = concl_of tr;
+ val (_$l$r) = Thm.concl_of tr;
val trs = if l aconv r then [] else [tr];
in rewrite_goals_tac ctxt trs th end))
--- a/src/HOL/Nat.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nat.thy Wed Mar 04 19:53:18 2015 +0100
@@ -1629,7 +1629,7 @@
declaration {* K Lin_Arith.setup *}
simproc_setup fast_arith_nat ("(m::nat) < n" | "(m::nat) <= n" | "(m::nat) = n") =
- {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (term_of ct) *}
+ {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
(* Because of this simproc, the arithmetic solver is really only
useful to detect inconsistencies among the premises for subgoals which are
*not* themselves (in)equalities, because the latter activate
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Wed Mar 04 19:53:18 2015 +0100
@@ -153,7 +153,7 @@
fun theorems_of thy =
filter (fn (name, th) =>
not (is_forbidden_theorem name) andalso
- (theory_of_thm th, thy) |> apply2 Context.theory_name |> op =)
+ (Thm.theory_of_thm th, thy) |> apply2 Context.theory_name |> op =)
(Global_Theory.all_thms_of thy true)
fun check_formulas tsp =
@@ -179,7 +179,7 @@
val _ = File.write path ""
fun check_theorem (name, th) =
let
- val t = th |> prop_of |> Type.legacy_freeze |> close_form
+ val t = th |> Thm.prop_of |> Type.legacy_freeze |> close_form
val neg_t = Logic.mk_implies (t, @{prop False})
val (nondef_ts, def_ts, _, _, _, _) =
TimeLimit.timeLimit preproc_timeout (preprocess_formulas hol_ctxt [])
--- a/src/HOL/Nominal/nominal_datatype.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML Wed Mar 04 19:53:18 2015 +0100
@@ -84,7 +84,7 @@
(*******************************)
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of Old_Datatype.distinct_lemma);
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (Thm.prems_of Old_Datatype.distinct_lemma);
(** simplification procedure for sorting permutations **)
@@ -108,9 +108,9 @@
val cp = cp_inst_of thy a b;
val dj = dj_thm_of thy b a;
val dj_cp' = [cp, dj] MRS dj_cp;
- val cert = SOME o cterm_of thy
+ val cert = SOME o Thm.cterm_of thy
in
- SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
+ SOME (mk_meta_eq (Drule.instantiate' [SOME (Thm.ctyp_of thy S)]
[cert t, cert r, cert s] dj_cp'))
end
else NONE
@@ -152,9 +152,11 @@
HOLogic.mk_prod (x, y) $ Const (@{const_name Nil}, pT)
end;
-fun mk_not_sym ths = maps (fn th => case prop_of th of
- _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) => [th, th RS not_sym]
- | _ => [th]) ths;
+fun mk_not_sym ths = maps (fn th =>
+ (case Thm.prop_of th of
+ _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) =>
+ [th, th RS not_sym]
+ | _ => [th])) ths;
fun fresh_const T U = Const (@{const_name Nominal.fresh}, T --> U --> HOLogic.boolT);
fun fresh_star_const T U =
@@ -775,11 +777,12 @@
fun dt_constr_defs ((((((_, (_, _, constrs)),
(_, (_, _, constrs'))), tname), T), T'), constr_syntax) (thy, defs, eqns, dist_lemmas) =
let
- val rep_const = cterm_of thy
+ val rep_const = Thm.cterm_of thy
(Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
val dist =
Drule.export_without_context
- (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] Old_Datatype.distinct_lemma);
+ (cterm_instantiate
+ [(Thm.cterm_of thy distinct_f, rep_const)] Old_Datatype.distinct_lemma);
val (thy', defs', eqns') = fold (make_constr_def tname T T')
(constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, [])
in
@@ -815,7 +818,7 @@
fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
let
- val _ $ (_ $ (Rep $ x)) = Logic.unvarify_global (prop_of th);
+ val _ $ (_ $ (Rep $ x)) = Logic.unvarify_global (Thm.prop_of th);
val Type ("fun", [T, U]) = fastype_of Rep;
val permT = mk_permT (Type (atom, []));
val pi = Free ("pi", permT);
@@ -1044,11 +1047,11 @@
full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps Rep_inverse_thms) 1,
etac mp 1, resolve_tac ctxt Rep_thms 1])]);
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma)));
val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
map (Free o apfst fst o dest_Var) Ps;
val indrule_lemma' = cterm_instantiate
- (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
+ (map (Thm.cterm_of thy8) Ps ~~ map (Thm.cterm_of thy8) frees) indrule_lemma;
val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
@@ -1254,17 +1257,18 @@
full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1,
REPEAT (etac conjE 1)])
[ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+ in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
fun fresh_fresh_inst thy a b =
let
val T = fastype_of a;
- val SOME th = find_first (fn th => case prop_of th of
+ val SOME th = find_first (fn th =>
+ (case Thm.prop_of th of
_ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
- | _ => false) perm_fresh_fresh
+ | _ => false)) perm_fresh_fresh
in
Drule.instantiate' []
- [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
+ [SOME (Thm.cterm_of thy a), NONE, SOME (Thm.cterm_of thy b)] th
end;
val fs_cp_sort =
@@ -1311,11 +1315,11 @@
SUBPROOF (fn {prems = iprems, params, concl,
context = context2, ...} =>
let
- val concl' = term_of concl;
+ val concl' = Thm.term_of concl;
val _ $ (_ $ _ $ u) = concl';
val U = fastype_of u;
val (xs, params') =
- chop (length cargs) (map (term_of o #2) params);
+ chop (length cargs) (map (Thm.term_of o #2) params);
val Ts = map fastype_of xs;
val cnstr = Const (cname, Ts ---> U);
val (pis, z) = split_last params';
@@ -1378,13 +1382,13 @@
val induct_aux' = Thm.instantiate ([],
map (fn (s, v as Var (_, T)) =>
- (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
+ (Thm.cterm_of thy9 v, Thm.cterm_of thy9 (Free (s, T))))
(pnames ~~ map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
+ (HOLogic.dest_Trueprop (Thm.concl_of induct_aux)))) @
map (fn (_, f) =>
let val f' = Logic.varify_global f
- in (cterm_of thy9 f',
- cterm_of thy9 (Const (@{const_name Nominal.supp}, fastype_of f')))
+ in (Thm.cterm_of thy9 f',
+ Thm.cterm_of thy9 (Const (@{const_name Nominal.supp}, fastype_of f')))
end) fresh_fs) induct_aux;
val induct = Goal.prove_global_future thy9 []
@@ -1547,8 +1551,8 @@
(augment_sort thy1 pt_cp_sort
(Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
(fn {context = ctxt, ...} => dtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), permT)),
- cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th) 1 THEN
+ [(Thm.cterm_of thy11 (Var (("pi", 0), permT)),
+ Thm.cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th) 1 THEN
NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths)
in (ths, ths') end) dt_atomTs);
@@ -1622,12 +1626,12 @@
fresh_prems) = chop (length finite_prems) prems;
val unique_prem' = unique_prem RS spec RS mp;
val unique = [unique_prem', unique_prem' RS sym] MRS trans;
- val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
+ val _ $ (_ $ (_ $ S $ _)) $ _ = Thm.prop_of supports_fresh;
val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
in EVERY
[rtac (Drule.cterm_instantiate
- [(cterm_of thy11 S,
- cterm_of thy11 (Const (@{const_name Nominal.supp},
+ [(Thm.cterm_of thy11 S,
+ Thm.cterm_of thy11 (Const (@{const_name Nominal.supp},
fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
supports_fresh) 1,
simp_tac (put_simpset HOL_basic_ss context addsimps
@@ -1638,8 +1642,9 @@
SUBPROOF (fn {prems = prems', params = [(_, a), (_, b)], ...} => EVERY
[cut_facts_tac [rec_prem] 1,
rtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
- cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
+ [(Thm.cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
+ Thm.cterm_of thy11
+ (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th) 1,
asm_simp_tac (put_simpset HOL_ss context addsimps
(prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
rtac rec_prem 1,
@@ -1666,7 +1671,7 @@
(rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
val induct_aux_rec = Drule.cterm_instantiate
- (map (apply2 (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
+ (map (apply2 (Thm.cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
(map (fn (aT, f) => (Logic.varify_global f, Abs ("z", HOLogic.unitT,
Const (@{const_name Nominal.supp}, fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
fresh_fs @
@@ -1694,7 +1699,7 @@
full_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_prod :: fresh_atm)) 1,
REPEAT (etac conjE 1)])
[ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+ in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
val finite_ctxt_prems = map (fn aT =>
HOLogic.mk_Trueprop
@@ -1747,9 +1752,9 @@
SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
let
val SOME prem = find_first (can (HOLogic.dest_eq o
- HOLogic.dest_Trueprop o prop_of)) prems';
- val _ $ (_ $ lhs $ rhs) = prop_of prem;
- val _ $ (_ $ lhs' $ rhs') = term_of concl;
+ HOLogic.dest_Trueprop o Thm.prop_of)) prems';
+ val _ $ (_ $ lhs $ rhs) = Thm.prop_of prem;
+ val _ $ (_ $ lhs' $ rhs') = Thm.term_of concl;
val rT = fastype_of lhs';
val (c, cargsl) = strip_comb lhs;
val cargsl' = partition_cargs idxs cargsl;
@@ -1758,17 +1763,20 @@
val cargsr' = partition_cargs idxs cargsr;
val boundsr = maps fst cargsr';
val (params1, _ :: params2) =
- chop (length params div 2) (map (term_of o #2) params);
+ chop (length params div 2) (map (Thm.term_of o #2) params);
val params' = params1 @ params2;
- val rec_prems = filter (fn th => case prop_of th of
- _ $ p => (case head_of p of
- Const (s, _) => member (op =) rec_set_names s
- | _ => false)
- | _ => false) prems';
- val fresh_prems = filter (fn th => case prop_of th of
+ val rec_prems = filter (fn th =>
+ (case Thm.prop_of th of
+ _ $ p =>
+ (case head_of p of
+ Const (s, _) => member (op =) rec_set_names s
+ | _ => false)
+ | _ => false)) prems';
+ val fresh_prems = filter (fn th =>
+ (case Thm.prop_of th of
_ $ (Const (@{const_name Nominal.fresh}, _) $ _ $ _) => true
| _ $ (Const (@{const_name Not}, _) $ _) => true
- | _ => false) prems';
+ | _ => false)) prems';
val Ts = map fastype_of boundsl;
val _ = warning "step 1: obtaining fresh names";
@@ -1842,7 +1850,7 @@
val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
val rec_prems' = map (fn th =>
let
- val _ $ (S $ x $ y) = prop_of th;
+ val _ $ (S $ x $ y) = Thm.prop_of th;
val Const (s, _) = head_of S;
val k = find_index (equal s) rec_set_names;
val pi = rpi1 @ pi2;
@@ -1853,8 +1861,8 @@
val l = find_index (equal T) dt_atomTs;
val th = nth (nth rec_equiv_thms' l) k;
val th' = Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), U)),
- cterm_of thy11 p)]) th;
+ [(Thm.cterm_of thy11 (Var (("pi", 0), U)),
+ Thm.cterm_of thy11 p)]) th;
in rtac th' 1 end;
val th' = Goal.prove context'' [] []
(HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
@@ -1868,15 +1876,17 @@
(put_simpset HOL_basic_ss context'' addsimps rpi1_pi2_eqs) th'
end) rec_prems2;
- val ihs = filter (fn th => case prop_of th of
- _ $ (Const (@{const_name All}, _) $ _) => true | _ => false) prems';
+ val ihs = filter (fn th =>
+ (case Thm.prop_of th of
+ _ $ (Const (@{const_name All}, _) $ _) => true
+ | _ => false)) prems';
(** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
val rec_eqns = map (fn (th, ih) =>
let
val th' = th RS (ih RS spec RS mp) RS sym;
- val _ $ (_ $ lhs $ rhs) = prop_of th';
+ val _ $ (_ $ lhs $ rhs) = Thm.prop_of th';
fun strip_perm (_ $ _ $ t) = strip_perm t
| strip_perm t = t;
in
@@ -1894,7 +1904,7 @@
maps (fn (rec_prem, ih) =>
let
val _ $ (S $ x $ (y as Free (_, T))) =
- prop_of rec_prem;
+ Thm.prop_of rec_prem;
val k = find_index (equal S) rec_sets;
val atoms = flat (map_filter (fn (bs, z) =>
if z = x then NONE else SOME bs) cargsl')
@@ -1981,7 +1991,7 @@
(fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
val _ = warning "final result";
- val final = Goal.prove context'' [] [] (term_of concl)
+ val final = Goal.prove context'' [] [] (Thm.term_of concl)
(fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
full_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh @
fresh_results @ fresh_results') 1);
--- a/src/HOL/Nominal/nominal_fresh_fun.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML Wed Mar 04 19:53:18 2015 +0100
@@ -12,22 +12,22 @@
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
fun gen_res_inst_tac_term ctxt instf tyinst tinst elim th i st =
let
- val thy = theory_of_thm st;
+ val thy = Thm.theory_of_thm st;
val cgoal = nth (cprems_of st) (i - 1);
- val {maxidx, ...} = rep_cterm cgoal;
+ val {maxidx, ...} = Thm.rep_cterm cgoal;
val j = maxidx + 1;
val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst;
- val ps = Logic.strip_params (term_of cgoal);
+ val ps = Logic.strip_params (Thm.term_of cgoal);
val Ts = map snd ps;
val tinst' = map (fn (t, u) =>
(head_of (Logic.incr_indexes (Ts, j) t),
fold_rev Term.abs ps u)) tinst;
val th' = instf
- (map (apply2 (ctyp_of thy)) tyinst')
- (map (apply2 (cterm_of thy)) tinst')
+ (map (apply2 (Thm.ctyp_of thy)) tyinst')
+ (map (apply2 (Thm.cterm_of thy)) tinst')
(Thm.lift_rule cgoal th)
in
- compose_tac ctxt (elim, th', nprems_of th) i st
+ compose_tac ctxt (elim, th', Thm.nprems_of th) i st
end handle General.Subscript => Seq.empty;
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
@@ -73,7 +73,7 @@
val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
(* find the variable we want to instantiate *)
- val x = hd (Misc_Legacy.term_vars (prop_of exists_fresh'));
+ val x = hd (Misc_Legacy.term_vars (Thm.prop_of exists_fresh'));
in
fn st =>
(cut_inst_tac_term' ctxt [(x,s)] exists_fresh' 1 THEN
@@ -131,7 +131,7 @@
val simp_ctxt =
ctxt addsimps (fresh_prod :: abs_fresh)
addsimps fresh_perm_app;
- val x = hd (tl (Misc_Legacy.term_vars (prop_of exI)));
+ val x = hd (tl (Misc_Legacy.term_vars (Thm.prop_of exI)));
val atom_name_opt = get_inner_fresh_fun goal;
val n = length (Logic.strip_params goal);
(* Here we rely on the fact that the variable introduced by generate_fresh_tac *)
@@ -145,18 +145,18 @@
val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
fun inst_fresh vars params i st =
- let val vars' = Misc_Legacy.term_vars (prop_of st);
- val thy = theory_of_thm st;
+ let val vars' = Misc_Legacy.term_vars (Thm.prop_of st);
+ val thy = Thm.theory_of_thm st;
in case subtract (op =) vars vars' of
[x] =>
- Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (fold_rev Term.abs params (Bound 0)))]) st)
+ Seq.single (Thm.instantiate ([], [(Thm.cterm_of thy x, Thm.cterm_of thy (fold_rev Term.abs params (Bound 0)))]) st)
| _ => error "fresh_fun_simp: Too many variables, please report."
end
in
((fn st =>
let
- val vars = Misc_Legacy.term_vars (prop_of st);
- val params = Logic.strip_params (nth (prems_of st) (i-1))
+ val vars = Misc_Legacy.term_vars (Thm.prop_of st);
+ val params = Logic.strip_params (nth (Thm.prems_of st) (i-1))
(* The tactics which solve the subgoals generated
by the conditionnal rewrite rule. *)
val post_rewrite_tacs =
--- a/src/HOL/Nominal/nominal_induct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_induct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -107,7 +107,7 @@
(CONJUNCTS (ALLGOALS
let
val adefs = nth_list atomized_defs (j - 1);
- val frees = fold (Term.add_frees o prop_of) adefs [];
+ val frees = fold (Term.add_frees o Thm.prop_of) adefs [];
val xs = nth_list fixings (j - 1);
val k = nth concls (j - 1) + more_consumes
in
--- a/src/HOL/Nominal/nominal_inductive.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML Wed Mar 04 19:53:18 2015 +0100
@@ -34,13 +34,13 @@
val perm_bool = mk_meta_eq @{thm perm_bool_def};
val perm_boolI = @{thm perm_boolI};
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
- (Drule.strip_imp_concl (cprop_of perm_boolI))));
+ (Drule.strip_imp_concl (Thm.cprop_of perm_boolI))));
fun mk_perm_bool pi th = th RS Drule.cterm_instantiate
[(perm_boolI_pi, pi)] perm_boolI;
fun mk_perm_bool_simproc names = Simplifier.simproc_global_i
- (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt =>
+ (Thm.theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt =>
fn Const (@{const_name Nominal.perm}, _) $ _ $ t =>
if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t))
then SOME perm_bool else NONE
@@ -126,7 +126,7 @@
fun map_thm ctxt f tac monos opt th =
let
- val prop = prop_of th;
+ val prop = Thm.prop_of th;
fun prove t =
Goal.prove ctxt [] [] t (fn _ =>
EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
@@ -141,7 +141,7 @@
eta_contract_cterm (Conjunction.mk_conjunction_balanced objs));
fun first_order_mrs ths th = ths MRS
- Thm.instantiate (first_order_matchs (cprems_of th) (map cprop_of ths)) th;
+ Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th;
fun prove_strong_ind s avoids ctxt =
let
@@ -159,7 +159,7 @@
commas_quote xs));
val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the
(Induct.lookup_inductP ctxt (hd names)))));
- val ([raw_induct'], ctxt') = Variable.import_terms false [prop_of raw_induct] ctxt;
+ val ([raw_induct'], ctxt') = Variable.import_terms false [Thm.prop_of raw_induct] ctxt;
val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
val ps = map (fst o snd) concls;
@@ -204,7 +204,7 @@
val fsT = TFree (fs_ctxt_tyname, ind_sort);
val inductive_forall_def' = Drule.instantiate'
- [SOME (ctyp_of thy fsT)] [] inductive_forall_def;
+ [SOME (Thm.ctyp_of thy fsT)] [] inductive_forall_def;
fun lift_pred' t (Free (s, T)) ts =
list_comb (Free (s, fsT --> T), t :: ts);
@@ -305,7 +305,7 @@
full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1,
REPEAT (etac conjE 1)])
[ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+ in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end;
fun mk_ind_proof ctxt' thss =
Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} =>
@@ -316,14 +316,14 @@
SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
let
val (params', (pis, z)) =
- chop (length params - length atomTs - 1) (map (term_of o #2) params) ||>
+ chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||>
split_last;
val bvars' = map
(fn (Bound i, T) => (nth params' (length params' - i), T)
| (t, T) => (t, T)) bvars;
val pi_bvars = map (fn (t, _) =>
fold_rev (NominalDatatype.mk_perm []) pis t) bvars';
- val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
+ val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl));
val (freshs1, freshs2, ctxt'') = fold
(obtain_fresh_name (ts @ pi_bvars))
(map snd bvars') ([], [], ctxt');
@@ -336,9 +336,9 @@
else pi2
end;
val pis'' = fold (concat_perm #> map) pis' pis;
- val env = Pattern.first_order_match thy (ihypt, prop_of ihyp)
+ val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp)
(Vartab.empty, Vartab.empty);
- val ihyp' = Thm.instantiate ([], map (apply2 (cterm_of thy))
+ val ihyp' = Thm.instantiate ([], map (apply2 (Thm.cterm_of thy))
(map (Envir.subst_term env) vs ~~
map (fold_rev (NominalDatatype.mk_perm [])
(rev pis' @ pis)) params' @ [z])) ihyp;
@@ -346,7 +346,7 @@
Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]
addsimprocs [NominalDatatype.perm_simproc])
(Simplifier.simplify (put_simpset eqvt_ss ctxt')
- (fold_rev (mk_perm_bool o cterm_of thy)
+ (fold_rev (mk_perm_bool o Thm.cterm_of thy)
(rev pis' @ pis) th));
val (gprems1, gprems2) = split_list
(map (fn (th, t) =>
@@ -360,7 +360,7 @@
val vc_compat_ths' = map (fn th =>
let
val th' = first_order_mrs gprems1 th;
- val (bop, lhs, rhs) = (case concl_of th' of
+ val (bop, lhs, rhs) = (case Thm.concl_of th' of
_ $ (fresh $ lhs $ rhs) =>
(fn t => fn u => fresh $ t $ u, lhs, rhs)
| _ $ (_ $ (_ $ lhs $ rhs)) =>
@@ -382,14 +382,14 @@
(HOLogic.mk_Trueprop (list_comb (P $ hd ts,
map (fold (NominalDatatype.mk_perm []) pis') (tl ts))))
(fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, rtac ihyp' 1,
- REPEAT_DETERM_N (nprems_of ihyp - length gprems)
+ REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems)
(simp_tac swap_simps_simpset 1),
REPEAT_DETERM_N (length gprems)
(simp_tac (put_simpset HOL_basic_ss ctxt''
addsimps [inductive_forall_def']
addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
resolve_tac ctxt'' gprems2 1)]));
- val final = Goal.prove ctxt'' [] [] (term_of concl)
+ val final = Goal.prove ctxt'' [] [] (Thm.term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt''
addsimps vc_compat_ths'' @ freshs2' @
perm_fresh_fresh @ fresh_atm) 1);
@@ -407,7 +407,7 @@
val cases_prems = map (fn ((name, avoids), rule) =>
let
- val ([rule'], ctxt') = Variable.import_terms false [prop_of rule] ctxt;
+ val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] ctxt;
val prem :: prems = Logic.strip_imp_prems rule';
val concl = Logic.strip_imp_concl rule'
in
@@ -472,7 +472,7 @@
rtac (first_order_mrs case_hyps case_hyp) 1
else
let
- val params' = map (term_of o #2 o nth (rev params)) is;
+ val params' = map (Thm.term_of o #2 o nth (rev params)) is;
val tab = params' ~~ map fst qs;
val (hyps1, hyps2) = chop (length args) case_hyps;
(* turns a = t and [x1 # t, ..., xn # t] *)
@@ -483,12 +483,12 @@
(map (fn th =>
let
val (cf, ct) =
- Thm.dest_comb (Thm.dest_arg (cprop_of th));
+ Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th));
val arg_cong' = Drule.instantiate'
- [SOME (ctyp_of_term ct)]
+ [SOME (Thm.ctyp_of_term ct)]
[NONE, SOME ct, SOME cf] (arg_cong RS iffD2);
val inst = Thm.first_order_match (ct,
- Thm.dest_arg (Thm.dest_arg (cprop_of th')))
+ Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th')))
in [th', th] MRS Thm.instantiate inst arg_cong'
end) ths1,
ths2)
@@ -505,17 +505,17 @@
val freshs2' = NominalDatatype.mk_not_sym freshs2;
val pis = map (NominalDatatype.perm_of_pair)
((freshs1 ~~ map fst qs) @ (params' ~~ freshs1));
- val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis);
- val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
+ val mk_pis = fold_rev mk_perm_bool (map (Thm.cterm_of thy) pis);
+ val obj = Thm.cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
(fn x as Free _ =>
if member (op =) args x then x
else (case AList.lookup op = tab x of
SOME y => y
| NONE => fold_rev (NominalDatatype.mk_perm []) pis x)
- | x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps));
+ | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps));
val inst = Thm.first_order_match (Thm.dest_arg
(Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
- val th = Goal.prove ctxt3 [] [] (term_of concl)
+ val th = Goal.prove ctxt3 [] [] (Thm.term_of concl)
(fn {context = ctxt4, ...} =>
rtac (Thm.instantiate inst case_hyp) 1 THEN
SUBPROOF (fn {prems = fresh_hyps, ...} =>
@@ -610,7 +610,7 @@
end;
val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp";
val (([t], [pi]), ctxt') = ctxt |>
- Variable.import_terms false [concl_of raw_induct] ||>>
+ Variable.import_terms false [Thm.concl_of raw_induct] ||>>
Variable.variant_fixes ["pi"];
val eqvt_simpset = put_simpset HOL_basic_ss ctxt' addsimps
(NominalThmDecls.get_eqvt_thms ctxt' @ perm_pi_simp) addsimprocs
@@ -621,7 +621,7 @@
fun eqvt_tac pi (intr, vs) st =
let
fun eqvt_err s =
- let val ([t], ctxt'') = Variable.import_terms true [prop_of intr] ctxt'
+ let val ([t], ctxt'') = Variable.import_terms true [Thm.prop_of intr] ctxt'
in error ("Could not prove equivariance for introduction rule\n" ^
Syntax.string_of_term ctxt'' t ^ "\n" ^ s)
end;
@@ -630,16 +630,16 @@
val prems' = map (fn th => the_default th (map_thm ctxt''
(split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems;
val prems'' = map (fn th => Simplifier.simplify eqvt_simpset
- (mk_perm_bool (cterm_of thy pi) th)) prems';
- val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
- map (cterm_of thy o NominalDatatype.mk_perm [] pi o term_of o #2) params)
+ (mk_perm_bool (Thm.cterm_of thy pi) th)) prems';
+ val intr' = Drule.cterm_instantiate (map (Thm.cterm_of thy) vs ~~
+ map (Thm.cterm_of thy o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params)
intr
in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac ctxt'' prems'')) 1
end) ctxt' 1 st
in
case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of
NONE => eqvt_err ("Rule does not match goal\n" ^
- Syntax.string_of_term ctxt' (hd (prems_of st)))
+ Syntax.string_of_term ctxt' (hd (Thm.prems_of st)))
| SOME (th, _) => Seq.single th
end;
val thss = map (fn atom =>
--- a/src/HOL/Nominal/nominal_inductive2.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML Wed Mar 04 19:53:18 2015 +0100
@@ -38,13 +38,13 @@
val perm_bool = mk_meta_eq @{thm perm_bool_def};
val perm_boolI = @{thm perm_boolI};
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb
- (Drule.strip_imp_concl (cprop_of perm_boolI))));
+ (Drule.strip_imp_concl (Thm.cprop_of perm_boolI))));
fun mk_perm_bool pi th = th RS Drule.cterm_instantiate
[(perm_boolI_pi, pi)] perm_boolI;
fun mk_perm_bool_simproc names = Simplifier.simproc_global_i
- (theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt =>
+ (Thm.theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn ctxt =>
fn Const (@{const_name Nominal.perm}, _) $ _ $ t =>
if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t))
then SOME perm_bool else NONE
@@ -131,7 +131,7 @@
fun map_thm ctxt f tac monos opt th =
let
- val prop = prop_of th;
+ val prop = Thm.prop_of th;
fun prove t =
Goal.prove ctxt [] [] t (fn _ =>
EVERY [cut_facts_tac [th] 1, etac rev_mp 1,
@@ -144,10 +144,10 @@
in (Logic.list_all (params, t), (rev vs, subst_bounds (vs, t))) end;
fun inst_params thy (vs, p) th cts =
- let val env = Pattern.first_order_match thy (p, prop_of th)
+ let val env = Pattern.first_order_match thy (p, Thm.prop_of th)
(Vartab.empty, Vartab.empty)
in Thm.instantiate ([],
- map (Envir.subst_term env #> cterm_of thy) vs ~~ cts) th
+ map (Envir.subst_term env #> Thm.cterm_of thy) vs ~~ cts) th
end;
fun prove_strong_ind s alt_name avoids ctxt =
@@ -168,7 +168,7 @@
(Induct.lookup_inductP ctxt (hd names)))));
val induct_cases' = if null induct_cases then replicate (length intrs) ""
else induct_cases;
- val ([raw_induct'], ctxt') = Variable.import_terms false [prop_of raw_induct] ctxt;
+ val ([raw_induct'], ctxt') = Variable.import_terms false [Thm.prop_of raw_induct] ctxt;
val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |>
HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb);
val ps = map (fst o snd) concls;
@@ -230,7 +230,7 @@
val fsT = TFree (fs_ctxt_tyname, ind_sort);
val inductive_forall_def' = Drule.instantiate'
- [SOME (ctyp_of thy fsT)] [] inductive_forall_def;
+ [SOME (Thm.ctyp_of thy fsT)] [] inductive_forall_def;
fun lift_pred' t (Free (s, T)) ts =
list_comb (Free (s, fsT --> T), t :: ts);
@@ -319,7 +319,7 @@
val fs_atom = Global_Theory.get_thm thy
("fs_" ^ Long_Name.base_name atom ^ "1");
val avoid_th = Drule.instantiate'
- [SOME (ctyp_of thy (fastype_of p))] [SOME (cterm_of thy p)]
+ [SOME (Thm.ctyp_of thy (fastype_of p))] [SOME (Thm.cterm_of thy p)]
([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding});
val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result
(fn ctxt' => EVERY
@@ -333,7 +333,7 @@
val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2);
val (pis1, pis2) = chop (length Ts1)
(map Bound (length pTs - 1 downto 0));
- val _ $ (f $ (_ $ pi $ l) $ r) = prop_of th2
+ val _ $ (f $ (_ $ pi $ l) $ r) = Thm.prop_of th2
val th2' =
Goal.prove ctxt' [] []
(Logic.list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop
@@ -343,7 +343,7 @@
full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps perm_set_forget) 1) |>
Simplifier.simplify (put_simpset eqvt_ss ctxt')
in
- (freshs @ [term_of cx],
+ (freshs @ [Thm.term_of cx],
ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt')
end;
@@ -358,12 +358,12 @@
let
val (cparams', (pis, z)) =
chop (length params - length atomTs - 1) (map #2 params) ||>
- (map term_of #> split_last);
- val params' = map term_of cparams'
+ (map Thm.term_of #> split_last);
+ val params' = map Thm.term_of cparams'
val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
val pi_sets = map (fn (t, _) =>
fold_rev (NominalDatatype.mk_perm []) pis t) sets';
- val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
+ val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl));
val gprems1 = map_filter (fn (th, t) =>
if null (preds_of ps t) then SOME th
else
@@ -374,7 +374,7 @@
let
val th' = gprems1 MRS inst_params thy p th cparams';
val (h, ts) =
- strip_comb (HOLogic.dest_Trueprop (concl_of th'))
+ strip_comb (HOLogic.dest_Trueprop (Thm.concl_of th'))
in
Goal.prove ctxt' [] []
(HOLogic.mk_Trueprop (list_comb (h,
@@ -399,12 +399,12 @@
val pis'' = fold_rev (concat_perm #> map) pis' pis;
val ihyp' = inst_params thy vs_ihypt ihyp
(map (fold_rev (NominalDatatype.mk_perm [])
- (pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]);
+ (pis' @ pis) #> Thm.cterm_of thy) params' @ [Thm.cterm_of thy z]);
fun mk_pi th =
Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]
addsimprocs [NominalDatatype.perm_simproc])
(Simplifier.simplify (put_simpset eqvt_ss ctxt')
- (fold_rev (mk_perm_bool o cterm_of thy)
+ (fold_rev (mk_perm_bool o Thm.cterm_of thy)
(pis' @ pis) th));
val gprems2 = map (fn (th, t) =>
if null (preds_of ps t) then mk_pi th
@@ -425,7 +425,7 @@
addsimps [inductive_forall_def']
addsimprocs [NominalDatatype.perm_simproc]) 1 THEN
resolve_tac ctxt'' gprems2 1)]));
- val final = Goal.prove ctxt'' [] [] (term_of concl)
+ val final = Goal.prove ctxt'' [] [] (Thm.term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt''
addsimps vc_compat_ths1' @ fresh_ths1 @
perm_freshs_freshs') 1);
--- a/src/HOL/Nominal/nominal_permeq.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_permeq.ML Wed Mar 04 19:53:18 2015 +0100
@@ -71,8 +71,8 @@
val supports_fresh_rule = @{thm "supports_fresh"};
(* pulls out dynamically a thm via the proof state *)
-fun dynamic_thms st name = Global_Theory.get_thms (theory_of_thm st) name;
-fun dynamic_thm st name = Global_Theory.get_thm (theory_of_thm st) name;
+fun dynamic_thms st name = Global_Theory.get_thms (Thm.theory_of_thm st) name;
+fun dynamic_thm st name = Global_Theory.get_thm (Thm.theory_of_thm st) name;
(* needed in the process of fully simplifying permutations *)
@@ -203,14 +203,14 @@
if pi1 <> pi2 then (* only apply the composition rule in this case *)
if T = U then
SOME (Drule.instantiate'
- [SOME (ctyp_of thy (fastype_of t))]
- [SOME (cterm_of thy pi1), SOME (cterm_of thy pi2), SOME (cterm_of thy t)]
+ [SOME (Thm.ctyp_of thy (fastype_of t))]
+ [SOME (Thm.cterm_of thy pi1), SOME (Thm.cterm_of thy pi2), SOME (Thm.cterm_of thy t)]
(mk_meta_eq ([Global_Theory.get_thm thy ("pt_"^tname'^"_inst"),
Global_Theory.get_thm thy ("at_"^tname'^"_inst")] MRS pt_perm_compose_aux)))
else
SOME (Drule.instantiate'
- [SOME (ctyp_of thy (fastype_of t))]
- [SOME (cterm_of thy pi1), SOME (cterm_of thy pi2), SOME (cterm_of thy t)]
+ [SOME (Thm.ctyp_of thy (fastype_of t))]
+ [SOME (Thm.cterm_of thy pi1), SOME (Thm.cterm_of thy pi2), SOME (Thm.cterm_of thy t)]
(mk_meta_eq (Global_Theory.get_thm thy ("cp_"^tname'^"_"^uname'^"_inst") RS
cp1_aux)))
else NONE
@@ -293,11 +293,11 @@
fun finite_guess_tac_i tactical ctxt i st =
let val goal = nth (cprems_of st) (i - 1)
in
- case Envir.eta_contract (Logic.strip_assums_concl (term_of goal)) of
+ case Envir.eta_contract (Logic.strip_assums_concl (Thm.term_of goal)) of
_ $ (Const (@{const_name finite}, _) $ (Const (@{const_name Nominal.supp}, T) $ x)) =>
let
val cert = Thm.cterm_of (Thm.theory_of_thm st);
- val ps = Logic.strip_params (term_of goal);
+ val ps = Logic.strip_params (Thm.term_of goal);
val Ts = rev (map snd ps);
val vs = collect_vars 0 x [];
val s = fold_rev (fn v => fn s =>
@@ -308,7 +308,7 @@
Term.range_type T) $ s);
val supports_rule' = Thm.lift_rule goal supports_rule;
val _ $ (_ $ S $ _) =
- Logic.strip_assums_concl (hd (prems_of supports_rule'));
+ Logic.strip_assums_concl (hd (Thm.prems_of supports_rule'));
val supports_rule'' = Drule.cterm_instantiate
[(cert (head_of S), cert s')] supports_rule'
val fin_supp = dynamic_thms st ("fin_supp")
@@ -337,11 +337,11 @@
val ctxt1 = ctxt addsimps [Thm.symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
val ctxt2 = ctxt addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
in
- case Logic.strip_assums_concl (term_of goal) of
+ case Logic.strip_assums_concl (Thm.term_of goal) of
_ $ (Const (@{const_name Nominal.fresh}, Type ("fun", [T, _])) $ _ $ t) =>
let
val cert = Thm.cterm_of (Thm.theory_of_thm st);
- val ps = Logic.strip_params (term_of goal);
+ val ps = Logic.strip_params (Thm.term_of goal);
val Ts = rev (map snd ps);
val vs = collect_vars 0 t [];
val s = fold_rev (fn v => fn s =>
@@ -352,7 +352,7 @@
(Const (@{const_name Nominal.supp}, fastype_of1 (Ts, s) --> HOLogic.mk_setT T) $ s);
val supports_fresh_rule' = Thm.lift_rule goal supports_fresh_rule;
val _ $ (_ $ S $ _) =
- Logic.strip_assums_concl (hd (prems_of supports_fresh_rule'));
+ Logic.strip_assums_concl (hd (Thm.prems_of supports_fresh_rule'));
val supports_fresh_rule'' = Drule.cterm_instantiate
[(cert (head_of S), cert s')] supports_fresh_rule'
in
--- a/src/HOL/Nominal/nominal_primrec.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML Wed Mar 04 19:53:18 2015 +0100
@@ -296,14 +296,14 @@
val rec_rewritess =
unflat (map (fn (_, (_, _, constrs)) => constrs) descr) rec_rewrites;
- val fvars = rec_rewrites |> hd |> concl_of |> HOLogic.dest_Trueprop |>
+ val fvars = rec_rewrites |> hd |> Thm.concl_of |> HOLogic.dest_Trueprop |>
HOLogic.dest_eq |> fst |> strip_comb |> snd |> take_prefix is_Var |> fst;
val (pvars, ctxtvars) = List.partition
(equal HOLogic.boolT o body_type o snd)
(subtract (op =)
- (Term.add_vars (concl_of (hd rec_rewrites)) [])
+ (Term.add_vars (Thm.concl_of (hd rec_rewrites)) [])
(fold_rev (Term.add_vars o Logic.strip_assums_concl)
- (prems_of (hd rec_rewrites)) []));
+ (Thm.prems_of (hd rec_rewrites)) []));
val cfs = defs' |> hd |> snd |> strip_comb |> snd |>
curry (List.take o swap) (length fvars) |> map cert;
val invs' = (case invs of
@@ -320,7 +320,7 @@
let
val (i, j, cargs) = mk_idx eq
val th = nth (nth rec_rewritess i) j;
- val cargs' = th |> concl_of |> HOLogic.dest_Trueprop |>
+ val cargs' = th |> Thm.concl_of |> HOLogic.dest_Trueprop |>
HOLogic.dest_eq |> fst |> strip_comb |> snd |> List.last |>
strip_comb |> snd
in (cargs, Logic.strip_imp_prems eq,
@@ -328,12 +328,12 @@
(map cert cargs' ~~ map (cert o Free) cargs)) th)
end) eqns';
- val prems = foldr1 (common_prefix op aconv) (map (prems_of o #3) rec_rewrites');
+ val prems = foldr1 (common_prefix op aconv) (map (Thm.prems_of o #3) rec_rewrites');
val cprems = map cert prems;
val asms = map Thm.assume cprems;
val premss = map (fn (cargs, eprems, eqn) =>
map (fn t => fold_rev (Logic.all o Free) cargs (Logic.list_implies (eprems, t)))
- (List.drop (prems_of eqn, length prems))) rec_rewrites';
+ (List.drop (Thm.prems_of eqn, length prems))) rec_rewrites';
val cpremss = map (map cert) premss;
val asmss = map (map Thm.assume) cpremss;
--- a/src/HOL/Nominal/nominal_thmdecls.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Nominal/nominal_thmdecls.ML Wed Mar 04 19:53:18 2015 +0100
@@ -132,7 +132,8 @@
fun eqvt_add_del_aux flag orig_thm context =
let
val thy = Context.theory_of context
- val thms_to_be_added = (case (prop_of orig_thm) of
+ val thms_to_be_added =
+ (case Thm.prop_of orig_thm of
(* case: eqvt-lemma is of the implicational form *)
(Const(@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop},_) $ hyp) $ (Const (@{const_name Trueprop},_) $ concl)) =>
let
--- a/src/HOL/Orderings.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Orderings.thy Wed Mar 04 19:53:18 2015 +0100
@@ -608,7 +608,7 @@
in
fun antisym_le_simproc ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
(le as Const (_, T)) $ r $ s =>
(let
val prems = Simplifier.prems_of ctxt;
@@ -627,7 +627,7 @@
| _ => NONE);
fun antisym_less_simproc ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
NotC $ ((less as Const(_,T)) $ r $ s) =>
(let
val prems = Simplifier.prems_of ctxt;
--- a/src/HOL/Probability/Giry_Monad.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Probability/Giry_Monad.thy Wed Mar 04 19:53:18 2015 +0100
@@ -184,7 +184,7 @@
fun subprob_cong thm ctxt = (
let
val thm' = Thm.transfer (Proof_Context.theory_of ctxt) thm
- val free = thm' |> concl_of |> HOLogic.dest_Trueprop |> dest_comb |> fst |>
+ val free = thm' |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_comb |> fst |>
dest_comb |> snd |> strip_abs_body |> head_of |> is_Free
in
if free then ([], Measurable.add_local_cong (thm' RS @{thm subprob_measurableD(2)}) ctxt)
--- a/src/HOL/Probability/measurable.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Probability/measurable.ML Wed Mar 04 19:53:18 2015 +0100
@@ -114,7 +114,7 @@
fun is_too_generic thm =
let
- val concl = concl_of thm
+ val concl = Thm.concl_of thm
val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl
in is_Var (head_of concl') end
@@ -126,19 +126,21 @@
if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f
fun nth_hol_goal thm i =
- HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (prems_of thm) (i - 1))))
+ HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (Thm.prems_of thm) (i - 1))))
fun dest_measurable_fun t =
(case t of
(Const (@{const_name "Set.member"}, _) $ f $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => f
| _ => raise (TERM ("not a measurability predicate", [t])))
-fun not_measurable_prop n thm = if length (prems_of thm) < n then false else
- (case nth_hol_goal thm n of
- (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "sets"}, _) $ _)) => false
- | (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => false
- | _ => true)
- handle TERM _ => true;
+fun not_measurable_prop n thm =
+ if length (Thm.prems_of thm) < n then false
+ else
+ (case nth_hol_goal thm n of
+ (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "sets"}, _) $ _)) => false
+ | (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => false
+ | _ => true)
+ handle TERM _ => true;
fun indep (Bound i) t b = i < b orelse t <= i
| indep (f $ t) top bot = indep f top bot andalso indep t top bot
@@ -201,7 +203,7 @@
fun measurable_tac ctxt facts =
let
fun debug_fact msg thm () =
- msg ^ " " ^ Pretty.str_of (Syntax.pretty_term ctxt (prop_of thm))
+ msg ^ " " ^ Pretty.str_of (Syntax.pretty_term ctxt (Thm.prop_of thm))
fun IF' c t i = COND (c i) (t i) no_tac
@@ -225,7 +227,7 @@
| is_sets_eq _ = false
val cong_thms = get_cong (Context.Proof ctxt) @
- filter (fn thm => concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts
+ filter (fn thm => Thm.concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts
fun sets_cong_tac i =
Subgoal.FOCUS (fn {context = ctxt', prems = prems, ...} => (
@@ -235,7 +237,7 @@
r_tac "cong intro" [elem_congI]
THEN' SOLVED' (fn i => REPEAT_DETERM (
((r_tac "cong solve" (cong_thms @ [@{thm refl}])
- ORELSE' IF' (fn i => fn thm => nprems_of thm > i)
+ ORELSE' IF' (fn i => fn thm => Thm.nprems_of thm > i)
(SOLVED' (asm_full_simp_tac ctxt''))) i)))
end) 1) ctxt i
THEN flexflex_tac ctxt
@@ -249,7 +251,7 @@
val f = dest_measurable_fun (HOLogic.dest_Trueprop t)
fun cert f = map (Option.map (f (Proof_Context.theory_of ctxt)))
fun inst (ts, Ts) =
- Drule.instantiate' (cert ctyp_of Ts) (cert cterm_of ts)
+ Drule.instantiate' (cert Thm.ctyp_of Ts) (cert Thm.cterm_of ts)
@{thm measurable_compose_countable}
in r_tac "split countable" (cnt_prefixes ctxt f |> map inst) i end
handle TERM _ => no_tac) 1)
@@ -267,7 +269,7 @@
fun simproc ctxt redex =
let
- val t = HOLogic.mk_Trueprop (term_of redex);
+ val t = HOLogic.mk_Trueprop (Thm.term_of redex);
fun tac {context = ctxt, prems = _ } =
SOLVE (measurable_tac ctxt (Simplifier.prems_of ctxt));
in try (fn () => Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}) () end;
--- a/src/HOL/Product_Type.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Product_Type.thy Wed Mar 04 19:53:18 2015 +0100
@@ -75,7 +75,7 @@
simproc_setup unit_eq ("x::unit") = {*
fn _ => fn _ => fn ct =>
- if HOLogic.is_unit (term_of ct) then NONE
+ if HOLogic.is_unit (Thm.term_of ct) then NONE
else SOME (mk_meta_eq @{thm unit_eq})
*}
@@ -579,8 +579,10 @@
| eta_proc _ _ = NONE;
end;
*}
-simproc_setup split_beta ("split f z") = {* fn _ => fn ctxt => fn ct => beta_proc ctxt (term_of ct) *}
-simproc_setup split_eta ("split f") = {* fn _ => fn ctxt => fn ct => eta_proc ctxt (term_of ct) *}
+simproc_setup split_beta ("split f z") =
+ {* fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct) *}
+simproc_setup split_eta ("split f") =
+ {* fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct) *}
lemmas split_beta [mono] = prod.case_eq_if
@@ -1309,7 +1311,7 @@
(* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
simproc_setup Collect_mem ("Collect t") = {*
fn _ => fn ctxt => fn ct =>
- (case term_of ct of
+ (case Thm.term_of ct of
S as Const (@{const_name Collect}, Type (@{type_name fun}, [_, T])) $ t =>
let val (u, _, ps) = HOLogic.strip_psplits t in
(case u of
--- a/src/HOL/Prolog/prolog.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Prolog/prolog.ML Wed Mar 04 19:53:18 2015 +0100
@@ -43,15 +43,15 @@
| _ (* atom *) => true;
val check_HOHH_tac1 = PRIMITIVE (fn thm =>
- if isG (concl_of thm) then thm else raise not_HOHH);
+ if isG (Thm.concl_of thm) then thm else raise not_HOHH);
val check_HOHH_tac2 = PRIMITIVE (fn thm =>
- if forall isG (prems_of thm) then thm else raise not_HOHH);
-fun check_HOHH thm = (if isD (concl_of thm) andalso forall isG (prems_of thm)
+ if forall isG (Thm.prems_of thm) then thm else raise not_HOHH);
+fun check_HOHH thm = (if isD (Thm.concl_of thm) andalso forall isG (Thm.prems_of thm)
then thm else raise not_HOHH);
fun atomizeD ctxt thm =
let
- fun at thm = case concl_of thm of
+ fun at thm = case Thm.concl_of thm of
_$(Const(@{const_name All} ,_)$Abs(s,_,_))=>
let val s' = if s="P" then "PP" else s
in at(thm RS (Rule_Insts.read_instantiate ctxt [(("x", 0), s')] [s'] spec)) end
--- a/src/HOL/Proofs/Lambda/WeakNorm.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Proofs/Lambda/WeakNorm.thy Wed Mar 04 19:53:18 2015 +0100
@@ -432,12 +432,12 @@
fun dummyf _ = error "dummy";
val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"};
-val (dB1, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct1));
-val ct1' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct1)) dB1);
+val (dB1, _) = @{code type_NF} (typing_of_term [] dummyf (Thm.term_of ct1));
+val ct1' = Thm.cterm_of @{theory} (term_of_dB [] (#T (Thm.rep_cterm ct1)) dB1);
val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"};
-val (dB2, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct2));
-val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2);
+val (dB2, _) = @{code type_NF} (typing_of_term [] dummyf (Thm.term_of ct2));
+val ct2' = Thm.cterm_of @{theory} (term_of_dB [] (#T (Thm.rep_cterm ct2)) dB2);
*}
end
--- a/src/HOL/Statespace/distinct_tree_prover.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Statespace/distinct_tree_prover.ML Wed Mar 04 19:53:18 2015 +0100
@@ -90,12 +90,12 @@
*)
fun instantiate instTs insts =
let
- val instTs' = map (fn (T, U) => (dest_TVar (typ_of T), typ_of U)) instTs;
+ val instTs' = map (fn (T, U) => (dest_TVar (Thm.typ_of T), Thm.typ_of U)) instTs;
fun substT x = (case AList.lookup (op =) instTs' x of NONE => TVar x | SOME T' => T');
fun mapT_and_recertify ct =
let
- val thy = theory_of_cterm ct;
- in (cterm_of thy (Term.map_types (Term.map_type_tvar substT) (term_of ct))) end;
+ val thy = Thm.theory_of_cterm ct;
+ in (Thm.cterm_of thy (Term.map_types (Term.map_type_tvar substT) (Thm.term_of ct))) end;
val insts' = map (apfst mapT_and_recertify) insts;
in Thm.instantiate (instTs, insts') end;
@@ -133,7 +133,8 @@
fun mtch (env as (tyinsts, insts)) =
fn (Var (ixn, T), ct) =>
(case AList.lookup (op =) insts ixn of
- NONE => (naive_typ_match (T, typ_of (ctyp_of_term ct)) tyinsts, (ixn, ct) :: insts)
+ NONE =>
+ (naive_typ_match (T, Thm.typ_of (Thm.ctyp_of_term ct)) tyinsts, (ixn, ct) :: insts)
| SOME _ => env)
| (f $ t, ct) =>
let val (cf, ct') = Thm.dest_comb ct;
@@ -144,14 +145,16 @@
fun discharge prems rule =
let
- val thy = theory_of_thm (hd prems);
+ val thy = Thm.theory_of_thm (hd prems);
val (tyinsts,insts) =
- fold naive_cterm_first_order_match (prems_of rule ~~ map cprop_of prems) ([], []);
+ fold naive_cterm_first_order_match (Thm.prems_of rule ~~ map Thm.cprop_of prems) ([], []);
val tyinsts' =
- map (fn (v, (S, U)) => (ctyp_of thy (TVar (v, S)), ctyp_of thy U)) tyinsts;
+ map (fn (v, (S, U)) =>
+ (Thm.ctyp_of thy (TVar (v, S)), Thm.ctyp_of thy U)) tyinsts;
val insts' =
- map (fn (idxn, ct) => (cterm_of thy (Var (idxn, typ_of (ctyp_of_term ct))), ct)) insts;
+ map (fn (idxn, ct) =>
+ (Thm.cterm_of thy (Var (idxn, Thm.typ_of (Thm.ctyp_of_term ct))), ct)) insts;
val rule' = Thm.instantiate (tyinsts', insts') rule;
in fold Thm.elim_implies prems rule' end;
@@ -160,7 +163,7 @@
val (l_in_set_root, x_in_set_root, r_in_set_root) =
let
val (Node_l_x_d, r) =
- cprop_of @{thm in_set_root}
+ Thm.cprop_of @{thm in_set_root}
|> Thm.dest_comb |> #2
|> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb;
val (Node_l, x) = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb;
@@ -170,7 +173,7 @@
val (x_in_set_left, r_in_set_left) =
let
val (Node_l_x_d, r) =
- cprop_of @{thm in_set_left}
+ Thm.cprop_of @{thm in_set_left}
|> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2
|> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb;
val x = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb |> #2;
@@ -179,7 +182,7 @@
val (x_in_set_right, l_in_set_right) =
let
val (Node_l, x) =
- cprop_of @{thm in_set_right}
+ Thm.cprop_of @{thm in_set_right}
|> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2
|> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2
|> Thm.dest_comb |> #1 |> Thm.dest_comb |> #1
@@ -207,25 +210,25 @@
val (ps, x_rest, y_rest) = split_common_prefix x_path y_path;
val dist_subtree_thm = dist_subtree ps dist_thm;
- val subtree = cprop_of dist_subtree_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
+ val subtree = Thm.cprop_of dist_subtree_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
val (_, [l, _, _, r]) = Drule.strip_comb subtree;
fun in_set ps tree =
let
val (_, [l, x, _, r]) = Drule.strip_comb tree;
- val xT = ctyp_of_term x;
+ val xT = Thm.ctyp_of_term x;
in
(case ps of
[] =>
instantiate
- [(ctyp_of_term x_in_set_root, xT)]
+ [(Thm.ctyp_of_term x_in_set_root, xT)]
[(l_in_set_root, l), (x_in_set_root, x), (r_in_set_root, r)] @{thm in_set_root}
| Left :: ps' =>
let
val in_set_l = in_set ps' l;
val in_set_left' =
instantiate
- [(ctyp_of_term x_in_set_left, xT)]
+ [(Thm.ctyp_of_term x_in_set_left, xT)]
[(x_in_set_left, x), (r_in_set_left, r)] @{thm in_set_left};
in discharge [in_set_l] in_set_left' end
| Right :: ps' =>
@@ -233,7 +236,7 @@
val in_set_r = in_set ps' r;
val in_set_right' =
instantiate
- [(ctyp_of_term x_in_set_right, xT)]
+ [(Thm.ctyp_of_term x_in_set_right, xT)]
[(x_in_set_right, x), (l_in_set_right, l)] @{thm in_set_right};
in discharge [in_set_r] in_set_right' end)
end;
@@ -286,29 +289,29 @@
@{thm subtract_Tip} |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2
|> Thm.dest_comb |> #2;
val [alpha] = ct |> Thm.ctyp_of_term |> Thm.dest_ctyp;
- in (alpha, #1 (dest_Var (term_of ct))) end;
+ in (alpha, #1 (dest_Var (Thm.term_of ct))) end;
in
fun subtractProver (Const (@{const_name Tip}, T)) ct dist_thm =
let
val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
- val thy = theory_of_cterm ct;
+ val thy = Thm.theory_of_cterm ct;
val [alphaI] = #2 (dest_Type T);
in
Thm.instantiate
- ([(alpha, ctyp_of thy alphaI)],
- [(cterm_of thy (Var (v, treeT alphaI)), ct')]) @{thm subtract_Tip}
+ ([(alpha, Thm.ctyp_of thy alphaI)],
+ [(Thm.cterm_of thy (Var (v, treeT alphaI)), ct')]) @{thm subtract_Tip}
end
| subtractProver (Const (@{const_name Node}, nT) $ l $ x $ d $ r) ct dist_thm =
let
val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
val (_, [cl, _, _, cr]) = Drule.strip_comb ct;
- val ps = the (find_tree x (term_of ct'));
+ val ps = the (find_tree x (Thm.term_of ct'));
val del_tree = deleteProver dist_thm ps;
val dist_thm' = discharge [del_tree, dist_thm] @{thm delete_Some_all_distinct};
- val sub_l = subtractProver (term_of cl) cl (dist_thm');
+ val sub_l = subtractProver (Thm.term_of cl) cl (dist_thm');
val sub_r =
- subtractProver (term_of cr) cr
+ subtractProver (Thm.term_of cr) cr
(discharge [sub_l, dist_thm'] @{thm subtract_Some_all_distinct_res});
in discharge [del_tree, sub_l, sub_r] @{thm subtract_Node} end;
@@ -317,7 +320,7 @@
fun distinct_implProver dist_thm ct =
let
val ctree = ct |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
- val sub = subtractProver (term_of ctree) ctree dist_thm;
+ val sub = subtractProver (Thm.term_of ctree) ctree dist_thm;
in @{thm subtract_Some_all_distinct} OF [sub, dist_thm] end;
fun get_fst_success f [] = NONE
@@ -329,8 +332,8 @@
fun neq_x_y ctxt x y name =
(let
val dist_thm = the (try (Proof_Context.get_thm ctxt) name);
- val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
- val tree = term_of ctree;
+ val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
+ val tree = Thm.term_of ctree;
val x_path = the (find_tree x tree);
val y_path = the (find_tree y tree);
val thm = distinctTreeProver dist_thm x_path y_path;
--- a/src/HOL/Statespace/state_fun.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Statespace/state_fun.ML Wed Mar 04 19:53:18 2015 +0100
@@ -56,14 +56,14 @@
let val thy = Proof_Context.theory_of ctxt in
(case t of (Const (@{const_name HOL.conj},_) $ P $ Q) =>
let
- val P_P' = Simplifier.rewrite ctxt (cterm_of thy P);
- val P' = P_P' |> prop_of |> Logic.dest_equals |> #2;
+ val P_P' = Simplifier.rewrite ctxt (Thm.cterm_of thy P);
+ val P' = P_P' |> Thm.prop_of |> Logic.dest_equals |> #2;
in
if isFalse P' then SOME (conj1_False OF [P_P'])
else
let
- val Q_Q' = Simplifier.rewrite ctxt (cterm_of thy Q);
- val Q' = Q_Q' |> prop_of |> Logic.dest_equals |> #2;
+ val Q_Q' = Simplifier.rewrite ctxt (Thm.cterm_of thy Q);
+ val Q' = Q_Q' |> Thm.prop_of |> Logic.dest_equals |> #2;
in
if isFalse Q' then SOME (conj2_False OF [Q_Q'])
else if isTrue P' andalso isTrue Q' then SOME (conj_True OF [P_P', Q_Q'])
@@ -141,12 +141,12 @@
| mk_upds s = (Var (("s", mi + 1), sT), mi + 2);
val ct =
- cterm_of thy (Const (@{const_name StateFun.lookup}, lT) $ destr $ n $ fst (mk_upds s));
+ Thm.cterm_of thy (Const (@{const_name StateFun.lookup}, lT) $ destr $ n $ fst (mk_upds s));
val basic_ss = #1 (Data.get (Context.Proof ctxt));
val ctxt' = ctxt |> Config.put simp_depth_limit 100 |> put_simpset basic_ss;
val thm = Simplifier.rewrite ctxt' ct;
in
- if (op aconv) (Logic.dest_equals (prop_of thm))
+ if (op aconv) (Logic.dest_equals (Thm.prop_of thm))
then NONE
else SOME thm
end
@@ -251,7 +251,7 @@
Goal.prove ctxt0 [] []
(Logic.list_all (vars, Logic.mk_equals (trm, trm')))
(fn _ => rtac meta_ext 1 THEN simp_tac ctxt1 1);
- val eq2 = Simplifier.asm_full_rewrite ctxt2 (Thm.dest_equals_rhs (cprop_of eq1));
+ val eq2 = Simplifier.asm_full_rewrite ctxt2 (Thm.dest_equals_rhs (Thm.cprop_of eq1));
in SOME (Thm.transitive eq1 eq2) end
| _ => NONE)
end
@@ -375,7 +375,7 @@
val ctxt = Context.proof_of context;
val (lookup_ss, ex_lookup_ss, simprocs_active) = Data.get context;
val (lookup_ss', ex_lookup_ss') =
- (case concl_of thm of
+ (case Thm.concl_of thm of
(_ $ ((Const (@{const_name Ex}, _) $ _))) =>
(lookup_ss, simpset_map ctxt (Simplifier.add_simp thm) ex_lookup_ss)
| _ =>
--- a/src/HOL/Statespace/state_space.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Statespace/state_space.ML Wed Mar 04 19:53:18 2015 +0100
@@ -190,8 +190,8 @@
fun neq_x_y ctxt x y =
(let
val dist_thm = the (get_dist_thm (Context.Proof ctxt) (#1 (dest_Free x)));
- val ctree = cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
- val tree = term_of ctree;
+ val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2;
+ val tree = Thm.term_of ctree;
val x_path = the (DistinctTreeProver.find_tree x tree);
val y_path = the (DistinctTreeProver.find_tree y tree);
val thm = DistinctTreeProver.distinctTreeProver dist_thm x_path y_path;
@@ -238,7 +238,7 @@
val dist_thm_name = distinct_compsN;
val dist_thm_full_name = dist_thm_name;
- fun comps_of_thm thm = prop_of thm
+ fun comps_of_thm thm = Thm.prop_of thm
|> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);
fun type_attr phi = Thm.declaration_attribute (fn thm => fn context =>
--- a/src/HOL/TLA/Action.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TLA/Action.thy Wed Mar 04 19:53:18 2015 +0100
@@ -116,7 +116,7 @@
val action_rewrite = int_rewrite
fun action_use ctxt th =
- case (concl_of th) of
+ case Thm.concl_of th of
Const _ $ (Const (@{const_name Valid}, _) $ _) =>
(flatten (action_unlift ctxt th) handle THM _ => th)
| _ => th;
--- a/src/HOL/TLA/Intensional.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TLA/Intensional.thy Wed Mar 04 19:53:18 2015 +0100
@@ -271,18 +271,18 @@
fun matchsome tha thb =
let fun hmatch 0 = raise THM("matchsome: no match", 0, [tha,thb])
| hmatch n = matchres tha n thb handle THM _ => hmatch (n-1)
- in hmatch (nprems_of thb) end
+ in hmatch (Thm.nprems_of thb) end
fun hflatten t =
- case (concl_of t) of
- Const _ $ (Const (@{const_name HOL.implies}, _) $ _ $ _) => hflatten (t RS mp)
- | _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t
+ case Thm.concl_of t of
+ Const _ $ (Const (@{const_name HOL.implies}, _) $ _ $ _) => hflatten (t RS mp)
+ | _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t
in
hflatten t
end
fun int_use ctxt th =
- case (concl_of th) of
+ case Thm.concl_of th of
Const _ $ (Const (@{const_name Valid}, _) $ _) =>
(flatten (int_unlift ctxt th) handle THM _ => th)
| _ => th
--- a/src/HOL/TLA/TLA.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TLA/TLA.thy Wed Mar 04 19:53:18 2015 +0100
@@ -126,7 +126,7 @@
val temp_rewrite = int_rewrite
fun temp_use ctxt th =
- case (concl_of th) of
+ case Thm.concl_of th of
Const _ $ (Const (@{const_name Intensional.Valid}, _) $ _) =>
((flatten (temp_unlift ctxt th)) handle THM _ => th)
| _ => th;
--- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -737,7 +737,7 @@
in
(*Instantiate the abstract rule based on the contents of the
required instance*)
- diff_and_instantiate ctxt abs_rule_thm (prop_of abs_rule_thm) rule_t
+ diff_and_instantiate ctxt abs_rule_thm (Thm.prop_of abs_rule_thm) rule_t
end
@@ -1089,7 +1089,7 @@
raise SKELETON
(*FIXME or classify it as follows:
[(Caboose,
- prop_of @{thm asm_rl}
+ Thm.prop_of @{thm asm_rl}
|> SOME,
SOME (@{thm asm_rl}, TRY (HEADGOAL atac)))]
*)
@@ -1097,17 +1097,17 @@
case hd skel of
Assumed =>
(hd skel,
- prop_of @{thm asm_rl}
+ Thm.prop_of @{thm asm_rl}
|> SOME,
SOME (@{thm asm_rl}, TRY (HEADGOAL atac))) :: rest memo ctxt
| Caboose =>
[(Caboose,
- prop_of @{thm asm_rl}
+ Thm.prop_of @{thm asm_rl}
|> SOME,
SOME (@{thm asm_rl}, TRY (HEADGOAL atac)))]
| Unconjoin =>
(hd skel,
- prop_of @{thm conjI}
+ Thm.prop_of @{thm conjI}
|> SOME,
SOME (@{thm conjI}, rtac @{thm conjI} 1)) :: rest memo ctxt
| Split (split_node, solved_node, antes) =>
@@ -1126,7 +1126,7 @@
simulate_split ctxt split_fmla minor_prems_assumps conclusion
in
(hd skel,
- prop_of split_thm
+ Thm.prop_of split_thm
|> SOME,
SOME (split_thm, rtac split_thm 1)) :: rest memo ctxt
end
@@ -1190,7 +1190,7 @@
|> Global_Theory.get_thm thy
in
(hd skel,
- prop_of (def_thm thy)
+ Thm.prop_of (def_thm thy)
|> SOME,
SOME (def_thm thy,
HEADGOAL (rtac_wrap def_thm))) :: rest memo ctxt
@@ -1206,7 +1206,7 @@
|> Global_Theory.get_thm thy
in
(hd skel,
- prop_of ax_thm
+ Thm.prop_of ax_thm
|> SOME,
SOME (ax_thm, rtac ax_thm 1)) :: rest memo ctxt
end
@@ -1256,11 +1256,11 @@
let
val thy = Proof_Context.theory_of ctxt
in
- (Synth_step "ccontr", prop_of @{thm ccontr} |> SOME,
+ (Synth_step "ccontr", Thm.prop_of @{thm ccontr} |> SOME,
SOME (@{thm ccontr}, rtac @{thm ccontr} 1)) ::
- (Synth_step "neg_eq_false", prop_of neg_eq_false |> SOME,
+ (Synth_step "neg_eq_false", Thm.prop_of neg_eq_false |> SOME,
SOME (neg_eq_false, dtac neg_eq_false 1)) ::
- (Synth_step "sas_if_needed_tac", prop_of @{thm asm_rl}(*FIXME *) |> SOME,
+ (Synth_step "sas_if_needed_tac", Thm.prop_of @{thm asm_rl} (*FIXME *) |> SOME,
SOME (sas_if_needed_tac ctxt prob_name)) ::
skel_to_naive_tactic_dbg prover_tac ctxt prob_name
(make_skeleton ctxt
--- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Wed Mar 04 19:53:18 2015 +0100
@@ -563,7 +563,7 @@
In addition to the abstract thm ("scheme_thm"), this function is
also supplied with the (sub)term of the abstract thm ("scheme_t")
we want to use in the diff, in case only part of "scheme_t"
- might be needed (not the whole "prop_of scheme_thm")*)
+ might be needed (not the whole "Thm.prop_of scheme_thm")*)
fun diff_and_instantiate ctxt scheme_thm scheme_t instance_t =
let
val thy = Proof_Context.theory_of ctxt
@@ -572,14 +572,14 @@
diff thy (scheme_t, instance_t)
(*valuation of type variables*)
- val typeval = map (apply2 (ctyp_of thy)) type_pairing
+ val typeval = map (apply2 (Thm.ctyp_of thy)) type_pairing
val typeval_env =
map (apfst dest_TVar) type_pairing
(*valuation of term variables*)
val termval =
map (apfst (type_devar typeval_env)) term_pairing
- |> map (apply2 (cterm_of thy))
+ |> map (apply2 (Thm.cterm_of thy))
in
Thm.instantiate (typeval, termval) scheme_thm
end
@@ -670,7 +670,7 @@
fun head_quantified_variable ctxt i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Wed Mar 04 19:53:18 2015 +0100
@@ -168,7 +168,7 @@
fun inst_parametermatch_tac ctxt thms i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -204,7 +204,7 @@
fun nominal_inst_parametermatch_tac ctxt thm i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -247,7 +247,7 @@
fun canonicalise_qtfr_order ctxt i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
in
@@ -649,7 +649,7 @@
val thy = Proof_Context.theory_of ctxt
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -689,7 +689,7 @@
val thy = Proof_Context.theory_of ctxt
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -764,7 +764,7 @@
val _ = @{assert} (arity > 0)
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -809,7 +809,7 @@
val thy = Proof_Context.theory_of ctxt
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -917,7 +917,7 @@
(* val contextualise = fold absdummy (map snd params) *)
val contextualise = fold absfree params
- val skolem_cts = map (contextualise #> cterm_of thy) skolem_terms
+ val skolem_cts = map (contextualise #> Thm.cterm_of thy) skolem_terms
(*now the instantiation code*)
@@ -937,7 +937,7 @@
fun make_var pre_var =
the_single pre_var
|> Var
- |> cterm_of thy
+ |> Thm.cterm_of thy
|> SOME
in
if null pre_var then NONE
@@ -1049,7 +1049,7 @@
fun find_dec_arity i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
in
@@ -1093,7 +1093,7 @@
fun breakdown_inference i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
in
@@ -1110,7 +1110,7 @@
val rule = extuni_dec_n ctxt arity
val rule_hyp =
- prop_of rule
+ Thm.prop_of rule
|> Logic.dest_implies
|> fst (*assuming that rule has single hypothesis*)
@@ -1208,7 +1208,7 @@
fun standard_cnf_type ctxt i : thm -> (TPTP_Reconstruct.formula_kind * int * bool) option = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
@@ -1681,7 +1681,7 @@
fun clause_consts_diff thm =
let
val t =
- prop_of thm
+ Thm.prop_of thm
|> Logic.dest_implies
|> fst
@@ -1707,7 +1707,7 @@
fun remove_redundant_quantification ctxt i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
in
@@ -1763,7 +1763,7 @@
fun remove_redundant_quantification_in_lit ctxt i = fn st =>
let
val gls =
- prop_of st
+ Thm.prop_of st
|> Logic.strip_horn
|> fst
in
@@ -2003,7 +2003,7 @@
thy
prob_name (#meta pannot) n
|> the
- |> (fn {inference_fmla, ...} => cterm_of thy inference_fmla)
+ |> (fn {inference_fmla, ...} => Thm.cterm_of thy inference_fmla)
|> oracle_iinterp
end
*}
--- a/src/HOL/TPTP/atp_theory_export.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TPTP/atp_theory_export.ML Wed Mar 04 19:53:18 2015 +0100
@@ -170,7 +170,7 @@
val problem =
facts
|> map (fn ((_, loc), th) =>
- ((Thm.get_name_hint th, loc), th |> prop_of |> mono ? monomorphize_term ctxt))
+ ((Thm.get_name_hint th, loc), th |> Thm.prop_of |> mono ? monomorphize_term ctxt))
|> generate_atp_problem ctxt format Axiom type_enc Exporter combsN false false true []
@{prop False}
|> #1 |> sort_wrt (heading_sort_key o fst)
--- a/src/HOL/TPTP/mash_export.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/TPTP/mash_export.ML Wed Mar 04 19:53:18 2015 +0100
@@ -58,7 +58,7 @@
| _ => ("", []))
fun has_thm_thy th thy =
- Context.theory_name thy = Context.theory_name (theory_of_thm th)
+ Context.theory_name thy = Context.theory_name (Thm.theory_of_thm th)
fun has_thys thys th = exists (has_thm_thy th) thys
@@ -98,7 +98,7 @@
fun do_fact ((_, stature), th) =
let
val name = nickname_of_thm th
- val feats = features_of ctxt (theory_of_thm th) stature [prop_of th]
+ val feats = features_of ctxt (Thm.theory_of_thm th) stature [Thm.prop_of th]
val s = encode_str name ^ ": " ^ encode_strs (sort string_ord feats) ^ "\n"
in
File.append path s
@@ -188,14 +188,14 @@
val _ = tracing ("Fact " ^ string_of_int j ^ ": " ^ name)
val isar_deps = isar_dependencies_of name_tabs th
val do_query = not (is_bad_query ctxt ho_atp step j th isar_deps)
- val goal_feats = features_of ctxt (theory_of_thm th) stature [prop_of th]
+ val goal_feats = features_of ctxt (Thm.theory_of_thm th) stature [Thm.prop_of th]
val access_facts = filter_accessible_from th new_facts @ old_facts
val (marker, deps) =
smart_dependencies_of ctxt params_opt access_facts name_tabs th isar_deps
fun extra_features_of (((_, stature), th), weight) =
- [prop_of th]
- |> features_of ctxt (theory_of_thm th) stature
+ [Thm.prop_of th]
+ |> features_of ctxt (Thm.theory_of_thm th) stature
|> map (rpair (weight * extra_feature_factor))
val query =
@@ -261,7 +261,7 @@
val suggs =
old_facts
|> filter_accessible_from th
- |> mepo_or_mash_suggested_facts ctxt (theory_of_thm th) params max_suggs hyp_ts
+ |> mepo_or_mash_suggested_facts ctxt (Thm.theory_of_thm th) params max_suggs hyp_ts
concl_t
|> map (nickname_of_thm o snd)
in
--- a/src/HOL/Tools/ATP/atp_problem_generate.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/ATP/atp_problem_generate.ML Wed Mar 04 19:53:18 2015 +0100
@@ -760,9 +760,9 @@
fun do_introduce_combinators ctxt Ts t =
let val thy = Proof_Context.theory_of ctxt in
t |> conceal_bounds Ts
- |> cterm_of thy
+ |> Thm.cterm_of thy
|> Meson_Clausify.introduce_combinators_in_cterm
- |> prop_of |> Logic.dest_equals |> snd
+ |> Thm.prop_of |> Logic.dest_equals |> snd
|> reveal_bounds Ts
end
(* A type variable of sort "{}" will make abstraction fail. *)
@@ -1211,7 +1211,7 @@
if exists_Const (member (op =) Meson.presimplified_consts o fst) t then
t |> Skip_Proof.make_thm (Proof_Context.theory_of ctxt)
|> Meson.presimplify ctxt
- |> prop_of
+ |> Thm.prop_of
else
t
@@ -1750,7 +1750,7 @@
I
else
ths ~~ (1 upto length ths)
- |> maps (dub_and_inst needs_sound o apfst (apsnd prop_of))
+ |> maps (dub_and_inst needs_sound o apfst (apsnd Thm.prop_of))
|> make_facts
|> union (op = o apply2 #iformula))
(if completish then completish_helper_table else helper_table)
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -580,7 +580,7 @@
| uncomb (Abs (s, T, t)) = Abs (s, T, uncomb t)
| uncomb (t as Const (x as (s, _))) =
(case AList.lookup (op =) combinator_table s of
- SOME thm => thm |> prop_of |> specialize_type thy x |> Logic.dest_equals |> snd
+ SOME thm => thm |> Thm.prop_of |> specialize_type thy x |> Logic.dest_equals |> snd
| NONE => t)
| uncomb t = t
in uncomb end
--- a/src/HOL/Tools/ATP/atp_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/ATP/atp_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -342,7 +342,7 @@
if exists_Const (fn (s, _) => s = @{const_name Not}) t then
t |> Skip_Proof.make_thm thy
|> Meson.cong_extensionalize_thm thy
- |> prop_of
+ |> Thm.prop_of
else
t
@@ -353,8 +353,8 @@
fun abs_extensionalize_term ctxt t =
if exists_Const is_fun_equality t then
let val thy = Proof_Context.theory_of ctxt in
- t |> cterm_of thy |> Meson.abs_extensionalize_conv ctxt
- |> prop_of |> Logic.dest_equals |> snd
+ t |> Thm.cterm_of thy |> Meson.abs_extensionalize_conv ctxt
+ |> Thm.prop_of |> Logic.dest_equals |> snd
end
else
t
@@ -405,7 +405,7 @@
fun strip_subgoal goal i ctxt =
let
val (t, (frees, params)) =
- Logic.goal_params (prop_of goal) i
+ Logic.goal_params (Thm.prop_of goal) i
||> (map dest_Free #> Variable.variant_frees ctxt [] #> `(map Free))
val hyp_ts = t |> Logic.strip_assums_hyp |> map (curry subst_bounds frees)
val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
--- a/src/HOL/Tools/ATP/atp_waldmeister.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/ATP/atp_waldmeister.ML Wed Mar 04 19:53:18 2015 +0100
@@ -516,7 +516,7 @@
if helper_lemmas_needed then
[(helpersN,
@{thms waldmeister_fol}
- |> map (fn th => (("", (Global, General)), preproc (prop_of th)))
+ |> map (fn th => (("", (Global, General)), preproc (Thm.prop_of th)))
|> map (fn ((s, _) ,t) => mk_formula helper_prefix s Axiom (eq_trm_to_atp thy t)))]
else
[]
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Wed Mar 04 19:53:18 2015 +0100
@@ -474,7 +474,7 @@
fun flip_rels lthy n thm =
let
- val Rs = Term.add_vars (prop_of thm) [];
+ val Rs = Term.add_vars (Thm.prop_of thm) [];
val Rs' = rev (drop (length Rs - n) Rs);
val cRs = map (fn f => (Proof_Context.cterm_of lthy (Var f), Proof_Context.cterm_of lthy (mk_flip f))) Rs';
in
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Wed Mar 04 19:53:18 2015 +0100
@@ -94,7 +94,7 @@
fun co_induct_inst_as_projs ctxt k thm =
let
- val fs = Term.add_vars (prop_of thm) []
+ val fs = Term.add_vars (Thm.prop_of thm) []
|> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
fun mk_cfp (f as (_, T)) =
(Proof_Context.cterm_of ctxt (Var f), Proof_Context.cterm_of ctxt (mk_proj T (num_binder_types T) k));
@@ -107,7 +107,7 @@
fun mk_case_transfer_tac ctxt rel_cases cases =
let
- val n = length (tl (prems_of rel_cases));
+ val n = length (tl (Thm.prems_of rel_cases));
in
REPEAT_DETERM (HEADGOAL (rtac rel_funI)) THEN
HEADGOAL (etac rel_cases) THEN
@@ -512,7 +512,7 @@
val assms_tac =
let val assms' = map (unfold_thms ctxt (@{thm id_bnf_def} :: ctor_defs)) assms in
fold (curry (op ORELSE')) (map (fn thm =>
- funpow (length (prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms')
+ funpow (length (Thm.prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms')
(etac FalseE)
end;
val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -469,7 +469,7 @@
val T = mk_tupleT_balanced tfrees;
in
@{thm asm_rl[of "ALL x. P x --> Q x" for P Q]}
- |> Drule.instantiate' [SOME (ctyp_of @{theory} T)] []
+ |> Drule.instantiate' [SOME (Thm.ctyp_of @{theory} T)] []
|> Raw_Simplifier.rewrite_goals_rule @{context} @{thms split_paired_All[THEN eq_reflection]}
|> (fn thm => impI RS funpow n (fn th => allI RS th) thm)
|> Thm.varifyT_global
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1343,7 +1343,7 @@
NONE => []
| SOME (exhaustive_code, rhs, raw_rhs, ctr_thms) =>
let
- val ms = map (Logic.count_prems o prop_of) ctr_thms;
+ val ms = map (Logic.count_prems o Thm.prop_of) ctr_thms;
val (raw_goal, goal) = (raw_rhs, rhs)
|> apply2 (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
#> curry Logic.list_all (map dest_Free fun_args));
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Wed Mar 04 19:53:18 2015 +0100
@@ -40,7 +40,7 @@
fun exhaust_inst_as_projs ctxt frees thm =
let
val num_frees = length frees;
- val fs = Term.add_vars (prop_of thm) [] |> filter (can dest_funT o snd);
+ val fs = Term.add_vars (Thm.prop_of thm) [] |> filter (can dest_funT o snd);
fun find s = find_index (curry (op =) s) frees;
fun mk_cfp (f as ((s, _), T)) =
(Proof_Context.cterm_of ctxt (Var f), Proof_Context.cterm_of ctxt (mk_proj T num_frees (find s)));
@@ -149,7 +149,7 @@
unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);
fun inst_split_eq ctxt split =
- (case prop_of split of
+ (case Thm.prop_of split of
@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
let
val s = Name.uu;
--- a/src/HOL/Tools/BNF/bnf_lfp_size.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML Wed Mar 04 19:53:18 2015 +0100
@@ -258,7 +258,7 @@
val size_thmss = map2 append size_simpss overloaded_size_simpss;
val size_gen_thmss = size_simpss
fun rhs_is_zero thm =
- let val Const (trueprop, _) $ (Const (eq, _) $ _ $ rhs) = prop_of thm in
+ let val Const (trueprop, _) $ (Const (eq, _) $ _ $ rhs) = Thm.prop_of thm in
trueprop = @{const_name Trueprop} andalso
eq = @{const_name HOL.eq} andalso
rhs = HOLogic.zero
--- a/src/HOL/Tools/BNF/bnf_tactics.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/BNF/bnf_tactics.ML Wed Mar 04 19:53:18 2015 +0100
@@ -36,7 +36,7 @@
(*stolen from Christian Urban's Cookbook (and adapted slightly)*)
fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
let
- val concl_pat = Drule.strip_imp_concl (cprop_of thm)
+ val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm)
val insts = Thm.first_order_match (concl_pat, concl)
in
rtac (Drule.instantiate_normalize insts thm) 1
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Wed Mar 04 19:53:18 2015 +0100
@@ -795,7 +795,7 @@
val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss;
fun has_undefined_rhs thm =
- (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
+ (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of
Const (@{const_name undefined}, _) => true
| _ => false);
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Wed Mar 04 19:53:18 2015 +0100
@@ -53,7 +53,7 @@
fun true_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term True});
fun false_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term False});
- val monomorphic_prop_of = prop_of o Thm.unvarify_global o Drule.zero_var_indexes;
+ val monomorphic_prop_of = Thm.prop_of o Thm.unvarify_global o Drule.zero_var_indexes;
fun massage_inject (tp $ (eqv $ (_ $ t $ u) $ rhs)) = tp $ (eqv $ mk_fcT_eq (t, u) $ rhs);
fun massage_distinct (tp $ (_ $ (_ $ t $ u))) = [tp $ false_eq (t, u), tp $ false_eq (u, t)];
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -178,7 +178,8 @@
fun subst_nonatomic_types [] = I
| subst_nonatomic_types inst = map_types (typ_subst_nonatomic inst);
-fun lhs_head_of thm = Term.head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))));
+fun lhs_head_of thm =
+ Term.head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))));
fun mk_predT Ts = Ts ---> HOLogic.boolT;
fun mk_pred1T T = mk_predT [T];
@@ -213,7 +214,7 @@
fun cterm_instantiate_pos cts thm =
let
val cert = Thm.cterm_of (Thm.theory_of_thm thm);
- val vars = Term.add_vars (prop_of thm) [];
+ val vars = Term.add_vars (Thm.prop_of thm) [];
val vars' = rev (drop (length vars - length cts) vars);
val ps = map_filter (fn (_, NONE) => NONE
| (var, SOME ct) => SOME (cert (Var var), ct)) (vars' ~~ cts);
--- a/src/HOL/Tools/Function/fun_cases.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/fun_cases.ML Wed Mar 04 19:53:18 2015 +0100
@@ -30,7 +30,7 @@
val info = Function.get_info ctxt f handle List.Empty => err ();
val {elims, pelims, is_partial, ...} = info;
val elims = if is_partial then pelims else the elims;
- val cprop = cterm_of thy prop;
+ val cprop = Thm.cterm_of thy prop;
fun mk_elim rl =
Thm.implies_intr cprop
(Tactic.rule_by_tactic ctxt (Inductive.mk_cases_tac ctxt) (Thm.assume cprop RS rl))
--- a/src/HOL/Tools/Function/function.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function.ML Wed Mar 04 19:53:18 2015 +0100
@@ -166,7 +166,7 @@
prepare_function do_print prep default_constraint fixspec eqns config lthy
in
lthy'
- |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
+ |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]]
|> Proof.refine (Method.primitive_text (K (K goal_state))) |> Seq.hd
end
--- a/src/HOL/Tools/Function/function_common.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function_common.ML Wed Mar 04 19:53:18 2015 +0100
@@ -339,11 +339,11 @@
fun import_function_data t ctxt =
let
val thy = Proof_Context.theory_of ctxt
- val ct = cterm_of thy t
+ val ct = Thm.cterm_of thy t
val inst_morph = lift_morphism thy o Thm.instantiate
fun match (trm, data) =
- SOME (transform_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
+ SOME (transform_function_data data (inst_morph (Thm.match (Thm.cterm_of thy trm, ct))))
handle Pattern.MATCH => NONE
in
get_first match (Item_Net.retrieve (get_functions ctxt) t)
--- a/src/HOL/Tools/Function/function_context_tree.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function_context_tree.ML Wed Mar 04 19:53:18 2015 +0100
@@ -87,7 +87,7 @@
fun cong_deps crule =
let
- val num_branches = map_index (apsnd branch_vars) (prems_of crule)
+ val num_branches = map_index (apsnd branch_vars) (Thm.prems_of crule)
in
Int_Graph.empty
|> fold (fn (i,_)=> Int_Graph.new_node (i,i)) num_branches
@@ -114,13 +114,13 @@
val thy = Proof_Context.theory_of ctxt
val tt' = Logic.mk_equals (Pattern.rewrite_term thy [(Free fvar, h)] [] t, t)
- val (c, subs) = (concl_of r, prems_of r)
+ val (c, subs) = (Thm.concl_of r, Thm.prems_of r)
val subst =
Pattern.match (Proof_Context.theory_of ctxt) (c, tt') (Vartab.empty, Vartab.empty)
val branches = map (mk_branch ctxt o Envir.beta_norm o Envir.subst_term subst) subs
val inst = map (fn v =>
- (cterm_of thy (Var v), cterm_of thy (Envir.subst_term subst (Var v))))
+ (Thm.cterm_of thy (Var v), Thm.cterm_of thy (Envir.subst_term subst (Var v))))
(Term.add_vars c [])
in
(cterm_instantiate inst r, dep, branches)
@@ -160,8 +160,8 @@
fun inst_tree thy fvar f tr =
let
- val cfvar = cterm_of thy fvar
- val cf = cterm_of thy f
+ val cfvar = Thm.cterm_of thy fvar
+ val cf = Thm.cterm_of thy f
fun inst_term t =
subst_bound(f, abstract_over (fvar, t))
@@ -174,7 +174,7 @@
| inst_tree_aux (RCall (t, str)) =
RCall (inst_term t, inst_tree_aux str)
and inst_branch ((fxs, assms), str) =
- ((fxs, map (Thm.assume o cterm_of thy o inst_term o prop_of) assms),
+ ((fxs, map (Thm.assume o Thm.cterm_of thy o inst_term o Thm.prop_of) assms),
inst_tree_aux str)
in
inst_tree_aux tr
@@ -185,15 +185,15 @@
fun compose (fs1, as1) (fs2, as2) = (fs1 @ fs2, as1 @ as2)
fun export_term (fixes, assumes) =
- fold_rev (curry Logic.mk_implies o prop_of) assumes
+ fold_rev (curry Logic.mk_implies o Thm.prop_of) assumes
#> fold_rev (Logic.all o Free) fixes
fun export_thm thy (fixes, assumes) =
- fold_rev (Thm.implies_intr o cprop_of) assumes
- #> fold_rev (Thm.forall_intr o cterm_of thy o Free) fixes
+ fold_rev (Thm.implies_intr o Thm.cprop_of) assumes
+ #> fold_rev (Thm.forall_intr o Thm.cterm_of thy o Free) fixes
fun import_thm thy (fixes, athms) =
- fold (Thm.forall_elim o cterm_of thy o Free) fixes
+ fold (Thm.forall_elim o Thm.cterm_of thy o Free) fixes
#> fold Thm.elim_implies athms
@@ -243,7 +243,7 @@
fun rewrite_by_tree ctxt h ih x tr =
let
val thy = Proof_Context.theory_of ctxt
- fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (cterm_of thy t), x)
+ fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (Thm.cterm_of thy t), x)
| rewrite_help fix h_as x (RCall (_ $ arg, st)) =
let
val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *)
@@ -251,10 +251,10 @@
val iha = import_thm thy (fix, h_as) ha (* (a', h a') : G *)
|> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner))))
(* (a, h a) : G *)
- val inst_ih = instantiate' [] [SOME (cterm_of thy arg)] ih
+ val inst_ih = instantiate' [] [SOME (Thm.cterm_of thy arg)] ih
val eq = Thm.implies_elim (Thm.implies_elim inst_ih lRi) iha (* h a = f a *)
- val h_a'_eq_h_a = Thm.combination (Thm.reflexive (cterm_of thy h)) inner
+ val h_a'_eq_h_a = Thm.combination (Thm.reflexive (Thm.cterm_of thy h)) inner
val h_a_eq_f_a = eq RS eq_reflection
val result = Thm.transitive h_a'_eq_h_a h_a_eq_f_a
in
--- a/src/HOL/Tools/Function/function_core.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function_core.ML Wed Mar 04 19:53:18 2015 +0100
@@ -152,10 +152,10 @@
val lhs = inst pre_lhs
val rhs = inst pre_rhs
- val cqs = map (cterm_of thy) qs
- val ags = map (Thm.assume o cterm_of thy) gs
+ val cqs = map (Thm.cterm_of thy) qs
+ val ags = map (Thm.assume o Thm.cterm_of thy) gs
- val case_hyp = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+ val case_hyp = Thm.assume (Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
in
ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
cqs = cqs, ags = ags, case_hyp = case_hyp }
@@ -201,7 +201,7 @@
val h_assum =
HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
|> fold_rev (Logic.all o Free) rcfix
|> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] []
|> abstract_over_list (rev qs)
@@ -235,7 +235,7 @@
val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
- val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
+ val lhsi_eq_lhsj = Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
in if j < i then
let
val compat = lookup_compat_thm j i cts
@@ -274,15 +274,15 @@
val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
val h_assums = map (fn RCInfo {h_assum, ...} =>
- Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+ Thm.assume (Thm.cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
val (eql, _) =
Function_Context_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree
val replace_lemma = (eql RS meta_eq_to_obj_eq)
- |> Thm.implies_intr (cprop_of case_hyp)
- |> fold_rev (Thm.implies_intr o cprop_of) h_assums
- |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> Thm.implies_intr (Thm.cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) h_assums
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
|> fold_rev Thm.forall_intr cqs
|> Thm.close_derivation
in
@@ -293,7 +293,8 @@
fun mk_uniqueness_clause thy globals compat_store clausei clausej RLj =
let
val Globals {h, y, x, fvar, ...} = globals
- val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
+ val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} =
+ clausei
val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
@@ -301,15 +302,17 @@
val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
val compat = get_compat_thm thy compat_store i j cctxi cctxj
- val Ghsj' = map (fn RCInfo {h_assum, ...} => Thm.assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+ val Ghsj' =
+ map (fn RCInfo {h_assum, ...} =>
+ Thm.assume (Thm.cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
val RLj_import = RLj
|> fold Thm.forall_elim cqsj'
|> fold Thm.elim_implies agsj'
|> fold Thm.elim_implies Ghsj'
- val y_eq_rhsj'h = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
- val lhsi_eq_lhsj' = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
+ val y_eq_rhsj'h = Thm.assume (Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+ val lhsi_eq_lhsj' = Thm.assume (Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
(* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
in
(trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
@@ -319,12 +322,12 @@
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
|> (fn it => trans OF [y_eq_rhsj'h, it])
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
- |> fold_rev (Thm.implies_intr o cprop_of) Ghsj'
- |> fold_rev (Thm.implies_intr o cprop_of) agsj'
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) Ghsj'
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) agsj'
(* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
- |> Thm.implies_intr (cprop_of y_eq_rhsj'h)
- |> Thm.implies_intr (cprop_of lhsi_eq_lhsj')
- |> fold_rev Thm.forall_intr (cterm_of thy h :: cqsj')
+ |> Thm.implies_intr (Thm.cprop_of y_eq_rhsj'h)
+ |> Thm.implies_intr (Thm.cprop_of lhsi_eq_lhsj')
+ |> fold_rev Thm.forall_intr (Thm.cterm_of thy h :: cqsj')
end
@@ -339,13 +342,13 @@
val ih_intro_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ih_intro
fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
- |> fold_rev (Thm.implies_intr o cprop_of) CCas
- |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
+ |> fold_rev (Thm.forall_intr o Thm.cterm_of thy o Free) RIvs
val existence = fold (curry op COMP o prep_RC) RCs lGI
- val P = cterm_of thy (mk_eq (y, rhsC))
- val G_lhs_y = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+ val P = Thm.cterm_of thy (mk_eq (y, rhsC))
+ val G_lhs_y = Thm.assume (Thm.cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
val unique_clauses =
map2 (mk_uniqueness_clause thy globals compat_store clausei) clauses rep_lemmas
@@ -356,31 +359,32 @@
|> Seq.list_of |> the_single
val uniqueness = G_cases
- |> Thm.forall_elim (cterm_of thy lhs)
- |> Thm.forall_elim (cterm_of thy y)
+ |> Thm.forall_elim (Thm.cterm_of thy lhs)
+ |> Thm.forall_elim (Thm.cterm_of thy y)
|> Thm.forall_elim P
|> Thm.elim_implies G_lhs_y
|> fold elim_implies_eta unique_clauses
- |> Thm.implies_intr (cprop_of G_lhs_y)
- |> Thm.forall_intr (cterm_of thy y)
+ |> Thm.implies_intr (Thm.cprop_of G_lhs_y)
+ |> Thm.forall_intr (Thm.cterm_of thy y)
- val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
+ val P2 = Thm.cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
val exactly_one =
- @{thm ex1I} |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
+ @{thm ex1I}
+ |> instantiate' [SOME (Thm.ctyp_of thy ranT)] [SOME P2, SOME (Thm.cterm_of thy rhsC)]
|> curry (op COMP) existence
|> curry (op COMP) uniqueness
|> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym])
- |> Thm.implies_intr (cprop_of case_hyp)
- |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> Thm.implies_intr (Thm.cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
|> fold_rev Thm.forall_intr cqs
val function_value =
existence
|> Thm.implies_intr ihyp
- |> Thm.implies_intr (cprop_of case_hyp)
- |> Thm.forall_intr (cterm_of thy x)
- |> Thm.forall_elim (cterm_of thy lhs)
+ |> Thm.implies_intr (Thm.cprop_of case_hyp)
+ |> Thm.forall_intr (Thm.cterm_of thy x)
+ |> Thm.forall_elim (Thm.cterm_of thy lhs)
|> curry (op RS) refl
in
(exactly_one, function_value)
@@ -397,12 +401,12 @@
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
- |> cterm_of thy
+ |> Thm.cterm_of thy
val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
- |> instantiate' [] [NONE, SOME (cterm_of thy h)]
+ |> instantiate' [] [NONE, SOME (Thm.cterm_of thy h)]
val _ = trace_msg (K "Proving Replacement lemmas...")
val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses
@@ -419,16 +423,17 @@
|> Thm.forall_elim_vars 0
|> fold (curry op COMP) ex1s
|> Thm.implies_intr (ihyp)
- |> Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
- |> Thm.forall_intr (cterm_of thy x)
+ |> Thm.implies_intr (Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+ |> Thm.forall_intr (Thm.cterm_of thy x)
|> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
- |> (fn it => fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+ |> (fn it =>
+ fold (Thm.forall_intr o Thm.cterm_of thy o Var) (Term.add_vars (Thm.prop_of it) []) it)
val goalstate = Conjunction.intr graph_is_function complete
|> Thm.close_derivation
|> Goal.protect 0
- |> fold_rev (Thm.implies_intr o cprop_of) compat
- |> Thm.implies_intr (cprop_of complete)
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) compat
+ |> Thm.implies_intr (Thm.cprop_of complete)
in
(goalstate, values)
end
@@ -458,7 +463,7 @@
let
val (qs, t) = dest_all_all orig_intro
val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T)
- val vars = Term.add_vars (prop_of thm) []
+ val vars = Term.add_vars (Thm.prop_of thm) []
val varmap = AList.lookup (op =) (frees ~~ map fst vars)
#> the_default ("", 0)
in
@@ -478,7 +483,7 @@
let
fun mk_h_assm (rcfix, rcassm, rcarg) =
HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
|> fold_rev (Logic.all o Free) rcfix
in
HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
@@ -511,7 +516,7 @@
fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
|> fold_rev (curry Logic.mk_implies) gs
|> fold_rev (Logic.all o Free) rcfix
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
@@ -549,7 +554,7 @@
let
fun inst_term t = subst_bound(f, abstract_over (fvar, t))
in
- (rcfix, map (Thm.assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+ (rcfix, map (Thm.assume o Thm.cterm_of thy o inst_term o Thm.prop_of) rcassm, inst_term rcarg)
end
@@ -563,19 +568,20 @@
val thy = Proof_Context.theory_of ctxt
val Globals {domT, z, ...} = globals
- fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
+ fun mk_psimp
+ (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
let
- val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
- val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
+ val lhs_acc = Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
+ val z_smaller = Thm.cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
in
((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
|> (fn it => it COMP graph_is_function)
|> Thm.implies_intr z_smaller
- |> Thm.forall_intr (cterm_of thy z)
+ |> Thm.forall_intr (Thm.cterm_of thy z)
|> (fn it => it COMP valthm)
|> Thm.implies_intr lhs_acc
|> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff])
- |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
end
in
@@ -594,23 +600,23 @@
val Globals {domT, x, z, a, P, D, ...} = globals
val acc_R = mk_acc domT R
- val x_D = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
- val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
+ val x_D = Thm.assume (Thm.cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+ val a_D = Thm.cterm_of thy (HOLogic.mk_Trueprop (D $ a))
- val D_subset = cterm_of thy (Logic.all x
+ val D_subset = Thm.cterm_of thy (Logic.all x
(Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
HOLogic.mk_Trueprop (D $ z)))))
- |> cterm_of thy
+ |> Thm.cterm_of thy
(* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
val ihyp = Logic.all_const domT $ Abs ("z", domT,
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
HOLogic.mk_Trueprop (P $ Bound 0)))
- |> cterm_of thy
+ |> Thm.cterm_of thy
val aihyp = Thm.assume ihyp
@@ -628,19 +634,19 @@
end
fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
- |> Thm.forall_elim (cterm_of thy rcarg)
+ |> Thm.forall_elim (Thm.cterm_of thy rcarg)
|> Thm.elim_implies llRI
- |> fold_rev (Thm.implies_intr o cprop_of) CCas
- |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
+ |> fold_rev (Thm.forall_intr o Thm.cterm_of thy o Free) RIvs
val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
val step = HOLogic.mk_Trueprop (P $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs
|> fold_rev (curry Logic.mk_implies) gs
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
+ |> Thm.cterm_of thy
val P_lhs = Thm.assume step
|> fold Thm.forall_elim cqs
@@ -648,12 +654,12 @@
|> fold Thm.elim_implies ags
|> fold Thm.elim_implies P_recs
- val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
+ val res = Thm.cterm_of thy (HOLogic.mk_Trueprop (P $ x))
|> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
|> Thm.symmetric (* P lhs == P x *)
|> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
- |> Thm.implies_intr (cprop_of case_hyp)
- |> fold_rev (Thm.implies_intr o cprop_of) ags
+ |> Thm.implies_intr (Thm.cprop_of case_hyp)
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
|> fold_rev Thm.forall_intr cqs
in
(res, step)
@@ -665,8 +671,8 @@
|> Thm.forall_elim_vars 0
|> fold (curry op COMP) cases (* P x *)
|> Thm.implies_intr ihyp
- |> Thm.implies_intr (cprop_of x_D)
- |> Thm.forall_intr (cterm_of thy x)
+ |> Thm.implies_intr (Thm.cprop_of x_D)
+ |> Thm.forall_intr (Thm.cterm_of thy x)
val subset_induct_rule =
acc_subset_induct
@@ -681,16 +687,16 @@
val simple_induct_rule =
subset_induct_rule
- |> Thm.forall_intr (cterm_of thy D)
- |> Thm.forall_elim (cterm_of thy acc_R)
+ |> Thm.forall_intr (Thm.cterm_of thy D)
+ |> Thm.forall_elim (Thm.cterm_of thy acc_R)
|> atac 1 |> Seq.hd
|> (curry op COMP) (acc_downward
- |> (instantiate' [SOME (ctyp_of thy domT)]
- (map (SOME o cterm_of thy) [R, x, z]))
- |> Thm.forall_intr (cterm_of thy z)
- |> Thm.forall_intr (cterm_of thy x))
- |> Thm.forall_intr (cterm_of thy a)
- |> Thm.forall_intr (cterm_of thy P)
+ |> (instantiate' [SOME (Thm.ctyp_of thy domT)]
+ (map (SOME o Thm.cterm_of thy) [R, x, z]))
+ |> Thm.forall_intr (Thm.cterm_of thy z)
+ |> Thm.forall_intr (Thm.cterm_of thy x))
+ |> Thm.forall_intr (Thm.cterm_of thy a)
+ |> Thm.forall_intr (Thm.cterm_of thy P)
in
simple_induct_rule
end
@@ -704,7 +710,7 @@
qglr = (oqs, _, _, _), ...} = clause
val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
|> fold_rev (curry Logic.mk_implies) gs
- |> cterm_of thy
+ |> Thm.cterm_of thy
in
Goal.init goal
|> (SINGLE (resolve_tac ctxt [accI] 1)) |> the
@@ -737,11 +743,11 @@
|> map (fn (ctxt, thm) => Function_Context_Tree.export_thm thy ctxt thm)
val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) used (* additional hyps *)
|> Function_Context_Tree.export_term (fixes, assumes)
- |> fold_rev (curry Logic.mk_implies o prop_of) ags
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
+ |> Thm.cterm_of thy
val thm = Thm.assume hyp
|> fold Thm.forall_elim cqs
@@ -750,7 +756,7 @@
|> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
- |> cterm_of thy |> Thm.assume
+ |> Thm.cterm_of thy |> Thm.assume
val acc = thm COMP ih_case
val z_acc_local = acc
@@ -786,42 +792,42 @@
val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
(domT --> domT --> boolT) --> boolT) $ R')
- |> cterm_of thy (* "wf R'" *)
+ |> Thm.cterm_of thy (* "wf R'" *)
(* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
val ihyp = Logic.all_const domT $ Abs ("z", domT,
Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
HOLogic.mk_Trueprop (acc_R $ Bound 0)))
- |> cterm_of thy
+ |> Thm.cterm_of thy
val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
- val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
+ val R_z_x = Thm.cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
val (hyps, cases) = fold (mk_nest_term_case ctxt globals R' ihyp_a) clauses ([], [])
in
R_cases
- |> Thm.forall_elim (cterm_of thy z)
- |> Thm.forall_elim (cterm_of thy x)
- |> Thm.forall_elim (cterm_of thy (acc_R $ z))
+ |> Thm.forall_elim (Thm.cterm_of thy z)
+ |> Thm.forall_elim (Thm.cterm_of thy x)
+ |> Thm.forall_elim (Thm.cterm_of thy (acc_R $ z))
|> curry op COMP (Thm.assume R_z_x)
|> fold_rev (curry op COMP) cases
|> Thm.implies_intr R_z_x
- |> Thm.forall_intr (cterm_of thy z)
+ |> Thm.forall_intr (Thm.cterm_of thy z)
|> (fn it => it COMP accI)
|> Thm.implies_intr ihyp
- |> Thm.forall_intr (cterm_of thy x)
+ |> Thm.forall_intr (Thm.cterm_of thy x)
|> (fn it => Drule.compose (it, 2, wf_induct_rule))
|> curry op RS (Thm.assume wfR')
|> forall_intr_vars
|> (fn it => it COMP allI)
|> fold Thm.implies_intr hyps
|> Thm.implies_intr wfR'
- |> Thm.forall_intr (cterm_of thy R')
- |> Thm.forall_elim (cterm_of thy (inrel_R))
+ |> Thm.forall_intr (Thm.cterm_of thy R')
+ |> Thm.forall_elim (Thm.cterm_of thy (inrel_R))
|> curry op RS wf_in_rel
|> full_simplify (put_simpset HOL_basic_ss ctxt addsimps [in_rel_def])
- |> Thm.forall_intr (cterm_of thy Rrel)
+ |> Thm.forall_intr (Thm.cterm_of thy Rrel)
end
@@ -892,7 +898,7 @@
fun mk_partial_rules provedgoal =
let
- val newthy = theory_of_thm provedgoal (*FIXME*)
+ val newthy = Thm.theory_of_thm provedgoal (*FIXME*)
val newctxt = Proof_Context.init_global newthy (*FIXME*)
val (graph_is_function, complete_thm) =
--- a/src/HOL/Tools/Function/function_elims.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function_elims.ML Wed Mar 04 19:53:18 2015 +0100
@@ -43,7 +43,7 @@
if Logic.occs (Free x, t) then raise Match else false
| _ => raise Match);
fun mk_eq thm =
- (if inspect (prop_of thm) then [thm RS eq_reflection]
+ (if inspect (Thm.prop_of thm) then [thm RS eq_reflection]
else [Thm.symmetric (thm RS eq_reflection)])
handle Match => [];
val simpset =
@@ -80,7 +80,7 @@
fun mk_partial_elim_rules ctxt result =
let
val thy = Proof_Context.theory_of ctxt;
- val cert = cterm_of thy;
+ val cert = Thm.cterm_of thy;
val FunctionResult {fs, R, dom, psimps, cases, ...} = result;
val n_fs = length fs;
@@ -98,14 +98,14 @@
val f_simps =
filter (fn r =>
- (prop_of r |> Logic.strip_assums_concl
+ (Thm.prop_of r |> Logic.strip_assums_concl
|> HOLogic.dest_Trueprop
|> dest_funprop |> fst |> fst) = f)
psimps;
val arity =
hd f_simps
- |> prop_of
+ |> Thm.prop_of
|> Logic.strip_assums_concl
|> HOLogic.dest_Trueprop
|> snd o fst o dest_funprop
--- a/src/HOL/Tools/Function/function_lib.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/function_lib.ML Wed Mar 04 19:53:18 2015 +0100
@@ -62,7 +62,7 @@
fun forall_intr_rename (n, cv) thm =
let
val allthm = Thm.forall_intr cv thm
- val (_ $ abs) = prop_of allthm
+ val (_ $ abs) = Thm.prop_of allthm
in Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy)) allthm end
@@ -95,17 +95,17 @@
fun regroup_conv neu cn ac is ct =
let
- val thy = theory_of_cterm ct
+ val thy = Thm.theory_of_cterm ct
val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *)
val mk = HOLogic.mk_binop cn
- val t = term_of ct
+ val t = Thm.term_of ct
val xs = dest_binop_list cn t
val js = subtract (op =) is (0 upto (length xs) - 1)
val ty = fastype_of t
in
Goal.prove_internal ctxt []
- (cterm_of thy
+ (Thm.cterm_of thy
(Logic.mk_equals (t,
if null is
then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
--- a/src/HOL/Tools/Function/induction_schema.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/induction_schema.ML Wed Mar 04 19:53:18 2015 +0100
@@ -46,8 +46,8 @@
(map meta (@{thm split_conv} :: @{thms sum.case}))
fun term_conv thy cv t =
- cv (cterm_of thy t)
- |> prop_of |> Logic.dest_equals |> snd
+ cv (Thm.cterm_of thy t)
+ |> Thm.prop_of |> Logic.dest_equals |> snd
fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T))
@@ -205,7 +205,7 @@
(IndScheme {T, cases=scases, branches}) =
let
val thy = Proof_Context.theory_of ctxt
- val cert = cterm_of thy
+ val cert = Thm.cterm_of thy
val n = length branches
val scases_idx = map_index I scases
@@ -274,14 +274,14 @@
|> Conv.fconv_rule (sum_prod_conv ctxt)
|> Conv.fconv_rule (ind_rulify ctxt)
|> (fn th => th COMP ipres) (* P rs *)
- |> fold_rev (Thm.implies_intr o cprop_of) cGas
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) cGas
|> fold_rev Thm.forall_intr cGvs
end
val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *)
val step = HOLogic.mk_Trueprop (list_comb (P, lhs))
- |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs
|> fold_rev (curry Logic.mk_implies) gs
|> fold_rev (Logic.all o Free) qs
|> cert
@@ -295,7 +295,7 @@
|> fold Thm.elim_implies ags
|> fold Thm.elim_implies P_recs (* P lhs *)
|> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *)
- |> fold_rev (Thm.implies_intr o cprop_of) (ags @ case_hyps)
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) (ags @ case_hyps)
|> fold_rev Thm.forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *)
in
(res, (cidx, step))
@@ -308,7 +308,7 @@
|> fold (Thm.forall_elim o cert) (fxs @ map Free ws)
|> fold Thm.elim_implies C_hyps
|> fold Thm.elim_implies cases (* P xs *)
- |> fold_rev (Thm.implies_intr o cprop_of) C_hyps
+ |> fold_rev (Thm.implies_intr o Thm.cprop_of) C_hyps
|> fold_rev (Thm.forall_intr o cert o Free) ws
val Pxs = cert (HOLogic.mk_Trueprop (P_comp $ x))
@@ -319,7 +319,7 @@
|> Seq.hd
|> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
|> Goal.finish ctxt
- |> Thm.implies_intr (cprop_of branch_hyp)
+ |> Thm.implies_intr (Thm.cprop_of branch_hyp)
|> fold_rev (Thm.forall_intr o cert) fxs
in
(Pxs, steps)
@@ -380,7 +380,7 @@
end
val res = Conjunction.intr_balanced (map_index project branches)
- |> fold_rev Thm.implies_intr (map cprop_of newgoals @ steps)
+ |> fold_rev Thm.implies_intr (map Thm.cprop_of newgoals @ steps)
|> Drule.generalize ([], [Rn])
val nbranches = length branches
--- a/src/HOL/Tools/Function/lexicographic_order.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/lexicographic_order.ML Wed Mar 04 19:53:18 2015 +0100
@@ -68,7 +68,7 @@
fun mk_cell thy solve_tac (vars, prems, lhs, rhs) mfun = Lazy.lazy (fn _ =>
let
- val goals = cterm_of thy o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs)
+ val goals = Thm.cterm_of thy o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs)
in
case try_proof (goals @{const_name Orderings.less}) solve_tac of
Solved thm => Less thm
@@ -76,7 +76,7 @@
(case try_proof (goals @{const_name Orderings.less_eq}) solve_tac of
Solved thm2 => LessEq (thm2, thm)
| Stuck thm2 =>
- if prems_of thm2 = [HOLogic.Trueprop $ @{term False}] then False thm2
+ if Thm.prems_of thm2 = [HOLogic.Trueprop $ @{term False}] then False thm2
else None (thm2, thm)
| _ => raise Match) (* FIXME *)
| _ => raise Match
@@ -176,7 +176,7 @@
fun lex_order_tac quiet ctxt solve_tac st = SUBGOAL (fn _ =>
let
val thy = Proof_Context.theory_of ctxt
- val ((_ $ (_ $ rel)) :: tl) = prems_of st
+ val ((_ $ (_ $ rel)) :: tl) = Thm.prems_of st
val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel))
@@ -201,7 +201,7 @@
else ()
in (* 4: proof reconstruction *)
- st |> (PRIMITIVE (cterm_instantiate [(cterm_of thy rel, cterm_of thy relation)])
+ st |> (PRIMITIVE (cterm_instantiate [(Thm.cterm_of thy rel, Thm.cterm_of thy relation)])
THEN (REPEAT (rtac @{thm "wf_mlex"} 1))
THEN (rtac @{thm "wf_empty"} 1)
THEN EVERY (map prove_row clean_table))
--- a/src/HOL/Tools/Function/measure_functions.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/measure_functions.ML Wed Mar 04 19:53:18 2015 +0100
@@ -21,7 +21,7 @@
DEPTH_SOLVE (resolve_tac ctxt (rev (Named_Theorems.get ctxt @{named_theorems measure_function})) 1)
(HOLogic.mk_Trueprop (mk_is_measure (Var (("f",0), T --> HOLogic.natT)))
|> Proof_Context.cterm_of ctxt |> Goal.init)
- |> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
+ |> Seq.map (Thm.prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
|> Seq.list_of
--- a/src/HOL/Tools/Function/mutual.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/mutual.ML Wed Mar 04 19:53:18 2015 +0100
@@ -160,13 +160,13 @@
val args = map inst pre_args
val rhs = inst pre_rhs
- val cqs = map (cterm_of thy) qs
- val ags = map (Thm.assume o cterm_of thy) gs
+ val cqs = map (Thm.cterm_of thy) qs
+ val ags = map (Thm.assume o Thm.cterm_of thy) gs
val import = fold Thm.forall_elim cqs
#> fold Thm.elim_implies ags
- val export = fold_rev (Thm.implies_intr o cprop_of) ags
+ val export = fold_rev (Thm.implies_intr o Thm.cprop_of) ags
#> fold_rev forall_intr_rename (oqnames ~~ cqs)
in
F ctxt (f, qs, gs, args, rhs) import export
@@ -199,7 +199,9 @@
fun mk_applied_form ctxt caTs thm =
let
val thy = Proof_Context.theory_of ctxt
- val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
+ val xs =
+ map_index (fn (i, T) =>
+ Thm.cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
in
fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm
|> Conv.fconv_rule (Thm.beta_conversion true)
--- a/src/HOL/Tools/Function/partial_function.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/partial_function.ML Wed Mar 04 19:53:18 2015 +0100
@@ -69,7 +69,7 @@
NONE => NONE
| SOME {case_thms, ...} =>
let
- val lhs = prop_of (hd case_thms)
+ val lhs = Thm.prop_of (hd case_thms)
|> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;
val arity = length (snd (strip_comb lhs));
val conv = funpow (length args - arity) Conv.fun_conv
@@ -112,7 +112,7 @@
fun cterm_instantiate' cts thm =
let
val thy = Thm.theory_of_thm thm;
- val vs = rev (Term.add_vars (prop_of thm) [])
+ val vs = rev (Term.add_vars (Thm.prop_of thm) [])
|> map (Thm.cterm_of thy o Var);
in
cterm_instantiate (zip_options vs cts) thm
--- a/src/HOL/Tools/Function/pat_completeness.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/pat_completeness.ML Wed Mar 04 19:53:18 2015 +0100
@@ -24,9 +24,9 @@
fun inst_free var inst = Thm.forall_elim inst o Thm.forall_intr var
fun inst_case_thm thy x P thm =
- let val [Pv, xv] = Term.add_vars (prop_of thm) []
+ let val [Pv, xv] = Term.add_vars (Thm.prop_of thm) []
in
- thm |> cterm_instantiate (map (apply2 (cterm_of thy)) [(Var xv, x), (Var Pv, P)])
+ thm |> cterm_instantiate (map (apply2 (Thm.cterm_of thy)) [(Var xv, x), (Var Pv, P)])
end
fun invent_vars constr i =
@@ -44,7 +44,7 @@
| filter_pats thy cons pvars (([], thm) :: pts) = raise Match
| filter_pats thy cons pvars (((pat as Free _) :: pats, thm) :: pts) =
let val inst = list_comb (cons, pvars)
- in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
+ in (inst :: pats, inst_free (Thm.cterm_of thy pat) (Thm.cterm_of thy inst) thm)
:: (filter_pats thy cons pvars pts)
end
| filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
@@ -58,7 +58,7 @@
let
val thy = Proof_Context.theory_of ctxt
val (_, subps) = strip_comb pat
- val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
+ val eqs = map (Thm.cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
val c_eq_pat = simplify (put_simpset HOL_basic_ss ctxt addsimps (map Thm.assume eqs)) c_assum
in
(subps @ pats,
@@ -72,13 +72,13 @@
let
val thy = Proof_Context.theory_of ctxt
val (avars, pvars, newidx) = invent_vars cons idx
- val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
+ val c_hyp = Thm.cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
val c_assum = Thm.assume c_hyp
val newpats = map (transform_pat ctxt avars c_assum) (filter_pats thy cons pvars pats)
in
o_alg ctxt P newidx (avars @ vs) newpats
|> Thm.implies_intr c_hyp
- |> fold_rev (Thm.forall_intr o cterm_of thy) avars
+ |> fold_rev (Thm.forall_intr o Thm.cterm_of thy) avars
end
| constr_case _ _ _ _ _ _ = raise Match
and o_alg _ P idx [] (([], Pthm) :: _) = Pthm
@@ -110,10 +110,10 @@
HOLogic.mk_Trueprop P
|> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)
|> fold_rev Logic.all qs
- |> cterm_of thy
+ |> Thm.cterm_of thy
val hyps = map2 mk_assum qss patss
- fun inst_hyps hyp qs = fold (Thm.forall_elim o cterm_of thy) qs (Thm.assume hyp)
+ fun inst_hyps hyp qs = fold (Thm.forall_elim o Thm.cterm_of thy) qs (Thm.assume hyp)
val assums = map2 inst_hyps hyps qss
in
o_alg ctxt P 2 xs (patss ~~ assums)
@@ -141,7 +141,7 @@
val patss = map (map snd) x_pats
val complete_thm = prove_completeness ctxt xs thesis qss patss
- |> fold_rev (Thm.forall_intr o cterm_of thy) vs
+ |> fold_rev (Thm.forall_intr o Thm.cterm_of thy) vs
in
PRIMITIVE (fn st => Drule.compose (complete_thm, i, st))
end
--- a/src/HOL/Tools/Function/relation.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/relation.ML Wed Mar 04 19:53:18 2015 +0100
@@ -16,11 +16,11 @@
(* tactic version *)
fun inst_state_tac inst st =
- case Term.add_vars (prop_of st) [] of
+ (case Term.add_vars (Thm.prop_of st) [] of
[v as (_, T)] =>
let val cert = Thm.cterm_of (Thm.theory_of_thm st);
in PRIMITIVE (Thm.instantiate ([], [(cert (Var v), cert (inst T))])) st end
- | _ => Seq.empty;
+ | _ => Seq.empty);
fun relation_tac ctxt rel i =
TRY (Function_Common.termination_rule_tac ctxt i)
@@ -30,7 +30,7 @@
(* version with type inference *)
fun inst_state_infer_tac ctxt rel st =
- case Term.add_vars (prop_of st) [] of
+ (case Term.add_vars (Thm.prop_of st) [] of
[v as (_, T)] =>
let
val cert = Proof_Context.cterm_of ctxt;
@@ -42,7 +42,7 @@
in
PRIMITIVE (Thm.instantiate ([], [(cert (Var v), rel')])) st
end
- | _ => Seq.empty;
+ | _ => Seq.empty);
fun relation_infer_tac ctxt rel i =
TRY (Function_Common.termination_rule_tac ctxt i)
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -158,7 +158,7 @@
(case Termination.get_descent D (nth cs cidx) m1 m2 of
SOME (Termination.Less thm) =>
if bStrict then thm
- else (thm COMP (Thm.lift_rule (cprop_of thm) @{thm less_imp_le}))
+ else (thm COMP (Thm.lift_rule (Thm.cprop_of thm) @{thm less_imp_le}))
| SOME (Termination.LessEq (thm, _)) =>
if not bStrict then thm
else raise Fail "get_desc_thm"
@@ -272,7 +272,7 @@
val level_mapping =
map_index pt_lev lev
|> Termination.mk_sumcases D (setT nat_pairT)
- |> cterm_of thy
+ |> Thm.cterm_of thy
in
PROFILE "Proof Reconstruction"
(CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv sl))) 1
--- a/src/HOL/Tools/Function/termination.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Function/termination.ML Wed Mar 04 19:53:18 2015 +0100
@@ -144,7 +144,7 @@
Const (@{const_abbrev Set.empty}, fastype_of c1))
|> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
in
- case Function_Lib.try_proof (cterm_of thy goal) chain_tac of
+ case Function_Lib.try_proof (Thm.cterm_of thy goal) chain_tac of
Function_Lib.Solved thm => SOME thm
| _ => NONE
end
@@ -169,7 +169,7 @@
fun mk_desc thy tac vs Gam l r m1 m2 =
let
fun try rel =
- try_proof (cterm_of thy
+ try_proof (Thm.cterm_of thy
(Logic.list_all (vs,
Logic.mk_implies (HOLogic.mk_Trueprop Gam,
HOLogic.mk_Trueprop (Const (rel, @{typ "nat => nat => bool"})
@@ -181,7 +181,7 @@
(case try @{const_name Orderings.less_eq} of
Solved thm2 => LessEq (thm2, thm)
| Stuck thm2 =>
- if prems_of thm2 = [HOLogic.Trueprop $ @{term False}]
+ if Thm.prems_of thm2 = [HOLogic.Trueprop $ @{term False}]
then False thm2 else None (thm2, thm)
| _ => raise Match) (* FIXME *)
| _ => raise Match
@@ -275,8 +275,8 @@
fun wf_union_tac ctxt st = SUBGOAL (fn _ =>
let
val thy = Proof_Context.theory_of ctxt
- val cert = cterm_of thy
- val ((_ $ (_ $ rel)) :: ineqs) = prems_of st
+ val cert = Thm.cterm_of thy
+ val ((_ $ (_ $ rel)) :: ineqs) = Thm.prems_of st
fun mk_compr ineq =
let
--- a/src/HOL/Tools/Lifting/lifting_bnf.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_bnf.ML Wed Mar 04 19:53:18 2015 +0100
@@ -20,7 +20,7 @@
fun Quotient_tac bnf ctxt i =
let
val rel_Grp = rel_Grp_of_bnf bnf
- fun get_lhs thm = thm |> concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
+ fun get_lhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
val vars = get_lhs rel_Grp |> strip_comb |> snd |> map_filter (try (strip_comb #> snd #> hd))
val UNIVs = map (fn var => HOLogic.mk_UNIV (var |> dest_Var |> snd |> dest_Type |> snd |> hd)) vars
val inst = map2 (curry (apply2 (Proof_Context.cterm_of ctxt))) vars UNIVs
--- a/src/HOL/Tools/Lifting/lifting_def.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_def.ML Wed Mar 04 19:53:18 2015 +0100
@@ -57,12 +57,12 @@
fun try_prove_reflexivity ctxt prop =
let
val thy = Proof_Context.theory_of ctxt
- val cprop = cterm_of thy prop
+ val cprop = Thm.cterm_of thy prop
val rule = @{thm ge_eq_refl}
- val concl_pat = Drule.strip_imp_concl (cprop_of rule)
+ val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule)
val insts = Thm.first_order_match (concl_pat, cprop)
val rule = Drule.instantiate_normalize insts rule
- val prop = hd (prems_of rule)
+ val prop = hd (Thm.prems_of rule)
in
case mono_eq_prover ctxt prop of
SOME thm => SOME (thm RS rule)
@@ -83,7 +83,7 @@
let
fun preprocess ctxt thm =
let
- val tm = (strip_args 2 o HOLogic.dest_Trueprop o concl_of) thm;
+ val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm;
val param_rel = (snd o dest_comb o fst o dest_comb) tm;
val thy = Proof_Context.theory_of ctxt;
val free_vars = Term.add_vars param_rel [];
@@ -99,25 +99,26 @@
end;
val subst = fold make_subst free_vars [];
- val csubst = map (apply2 (cterm_of thy)) subst;
+ val csubst = map (apply2 (Thm.cterm_of thy)) subst;
val inst_thm = Drule.cterm_instantiate csubst thm;
in
Conv.fconv_rule
- ((Conv.concl_conv (nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv)
+ ((Conv.concl_conv (Thm.nprems_of inst_thm) o
+ HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv)
(Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm
end
fun inst_relcomppI thy ant1 ant2 =
let
- val t1 = (HOLogic.dest_Trueprop o concl_of) ant1
- val t2 = (HOLogic.dest_Trueprop o prop_of) ant2
- val fun1 = cterm_of thy (strip_args 2 t1)
- val args1 = map (cterm_of thy) (get_args 2 t1)
- val fun2 = cterm_of thy (strip_args 2 t2)
- val args2 = map (cterm_of thy) (get_args 1 t2)
+ val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1
+ val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2
+ val fun1 = Thm.cterm_of thy (strip_args 2 t1)
+ val args1 = map (Thm.cterm_of thy) (get_args 2 t1)
+ val fun2 = Thm.cterm_of thy (strip_args 2 t2)
+ val args2 = map (Thm.cterm_of thy) (get_args 1 t2)
val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI}
- val vars = (rev (Term.add_vars (prop_of relcomppI) []))
- val subst = map (apfst ((cterm_of thy) o Var)) (vars ~~ ([fun1] @ args1 @ [fun2] @ args2))
+ val vars = (rev (Term.add_vars (Thm.prop_of relcomppI) []))
+ val subst = map (apfst (Thm.cterm_of thy o Var)) (vars ~~ ([fun1] @ args1 @ [fun2] @ args2))
in
Drule.cterm_instantiate subst relcomppI
end
@@ -126,11 +127,12 @@
let
val thy = Proof_Context.theory_of ctxt
fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT)
- val rel = (Thm.dest_fun2 o Thm.dest_arg o cprop_of) thm
- val typ = (typ_of o ctyp_of_term) rel
- val POS_const = cterm_of thy (mk_POS typ)
- val var = cterm_of thy (Var (("X", #maxidx (rep_cterm (rel)) + 1), typ))
- val goal = Thm.apply (cterm_of thy HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var)
+ val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm
+ val typ = (Thm.typ_of o Thm.ctyp_of_term) rel
+ val POS_const = Thm.cterm_of thy (mk_POS typ)
+ val var = Thm.cterm_of thy (Var (("X", #maxidx (Thm.rep_cterm (rel)) + 1), typ))
+ val goal =
+ Thm.apply (Thm.cterm_of thy HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var)
in
[Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply}
end
@@ -208,13 +210,13 @@
fun unabs_def ctxt def =
let
- val (_, rhs) = Thm.dest_equals (cprop_of def)
+ val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
fun dest_abs (Abs (var_name, T, _)) = (var_name, T)
| dest_abs tm = raise TERM("get_abs_var",[tm])
- val (var_name, T) = dest_abs (term_of rhs)
+ val (var_name, T) = dest_abs (Thm.term_of rhs)
val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt
val thy = Proof_Context.theory_of ctxt'
- val refl_thm = Thm.reflexive (cterm_of thy (Free (hd new_var_names, T)))
+ val refl_thm = Thm.reflexive (Thm.cterm_of thy (Free (hd new_var_names, T)))
in
Thm.combination def refl_thm |>
singleton (Proof_Context.export ctxt' ctxt)
@@ -222,8 +224,8 @@
fun unabs_all_def ctxt def =
let
- val (_, rhs) = Thm.dest_equals (cprop_of def)
- val xs = strip_abs_vars (term_of rhs)
+ val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
+ val xs = strip_abs_vars (Thm.term_of rhs)
in
fold (K (unabs_def ctxt)) xs def
end
@@ -295,11 +297,11 @@
val rel_fun = prove_rel ctxt rsp_thm (rty, qty)
val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq}
in
- case mono_eq_prover ctxt (hd(prems_of rep_abs_thm)) of
+ case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of
SOME mono_eq_thm =>
let
val rep_abs_eq = mono_eq_thm RS rep_abs_thm
- val rep = (cterm_of thy o quot_thm_rep) quot_thm
+ val rep = (Thm.cterm_of thy o quot_thm_rep) quot_thm
val rep_refl = Thm.reflexive rep RS @{thm meta_eq_to_obj_eq}
val repped_eq = [rep_refl, unabs_def RS @{thm meta_eq_to_obj_eq}] MRSL @{thm cong}
val code_cert = [repped_eq, rep_abs_eq] MRSL trans
@@ -323,11 +325,11 @@
val rep_abs_folded_unmapped_thm =
let
val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq}
- val ctm = Thm.dest_equals_lhs (cprop_of rep_id)
+ val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id)
val unfolded_maps_eq = unfold_fun_maps ctm
val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap}
val prems_pat = (hd o Drule.cprems_of) t1
- val insts = Thm.first_order_match (prems_pat, cprop_of unfolded_maps_eq)
+ val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq)
in
unfolded_maps_eq RS (Drule.instantiate_normalize insts t1)
end
@@ -354,7 +356,7 @@
| no_abstr (Const (name, _)) = not (Code.is_abstr thy name)
| no_abstr _ = true
fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true)
- andalso no_abstr (prop_of eqn)
+ andalso no_abstr (Thm.prop_of eqn)
fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq)
in
@@ -378,7 +380,7 @@
local
fun encode_code_eq thy abs_eq opt_rep_eq (rty, qty) =
let
- fun mk_type typ = typ |> Logic.mk_type |> cterm_of thy |> Drule.mk_term
+ fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of thy |> Drule.mk_term
in
Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty]
end
@@ -386,12 +388,12 @@
exception DECODE
fun decode_code_eq thm =
- if nprems_of thm > 0 then raise DECODE
+ if Thm.nprems_of thm > 0 then raise DECODE
else
let
val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm
val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq
- fun dest_type typ = typ |> Drule.dest_term |> term_of |> Logic.dest_type
+ fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type
in
(abs_eq, opt_rep_eq, (dest_type rty, dest_type qty))
end
@@ -508,7 +510,7 @@
map_interrupt prove assms
end
- fun cconl_of thm = Drule.strip_imp_concl (cprop_of thm)
+ fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm)
fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm))
fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm))
val rule1 = Thm.incr_indexes (#maxidx (Thm.rep_cterm ct) + 1) rule;
@@ -622,7 +624,7 @@
val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs)
|> Proof_Context.cterm_of lthy
|> cr_to_pcr_conv
- |> ` concl_of
+ |> `Thm.concl_of
|>> Logic.dest_equals
|>> snd
val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2
@@ -637,7 +639,7 @@
| NONE =>
let
val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy
- val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq)
+ val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
val readable_rsp_tm_tnames = rename_to_tnames lthy readable_rsp_tm
fun after_qed' thm_list lthy =
--- a/src/HOL/Tools/Lifting/lifting_info.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_info.ML Wed Mar 04 19:53:18 2015 +0100
@@ -160,7 +160,7 @@
end
val ((_, [rel_quot_thm_fixed]), ctxt') = Variable.importT [rel_quot_thm] ctxt
- val rel_quot_thm_prop = prop_of rel_quot_thm_fixed
+ val rel_quot_thm_prop = Thm.prop_of rel_quot_thm_fixed
val rel_quot_thm_concl = Logic.strip_imp_concl rel_quot_thm_prop
val rel_quot_thm_prems = Logic.strip_imp_prems rel_quot_thm_prop;
val concl_absT = quot_term_absT ctxt' rel_quot_thm_concl
@@ -184,7 +184,7 @@
fun add_quot_map rel_quot_thm ctxt =
let
val _ = Context.cases (K ()) (quot_map_thm_sanity_check rel_quot_thm) ctxt
- val rel_quot_thm_concl = (Logic.strip_imp_concl o prop_of) rel_quot_thm
+ val rel_quot_thm_concl = (Logic.strip_imp_concl o Thm.prop_of) rel_quot_thm
val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) rel_quot_thm_concl
val relatorT_name = (fst o dest_Type o fst o dest_funT o fastype_of) abs
val minfo = {rel_quot_thm = rel_quot_thm}
@@ -204,7 +204,7 @@
[Pretty.str "type:",
Pretty.str ty_name,
Pretty.str "quot. theorem:",
- Syntax.pretty_term ctxt (prop_of rel_quot_thm)])
+ Syntax.pretty_term ctxt (Thm.prop_of rel_quot_thm)])
in
map prt_map (Symtab.dest (get_quot_maps ctxt))
|> Pretty.big_list "maps for type constructors:"
@@ -243,11 +243,11 @@
[Pretty.str "type:",
Pretty.str qty_name,
Pretty.str "quot. thm:",
- Syntax.pretty_term ctxt (prop_of quot_thm),
+ Syntax.pretty_term ctxt (Thm.prop_of quot_thm),
Pretty.str "pcrel_def thm:",
- option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o prop_of o #pcrel_def) pcr_info,
+ option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o Thm.prop_of o #pcrel_def) pcr_info,
Pretty.str "pcr_cr_eq thm:",
- option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o prop_of o #pcr_cr_eq) pcr_info])
+ option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o Thm.prop_of o #pcr_cr_eq) pcr_info])
in
map prt_quot (Symtab.dest (get_quotients ctxt))
|> Pretty.big_list "quotients:"
@@ -304,13 +304,13 @@
fun introduce_polarities rule =
let
val dest_less_eq = HOLogic.dest_bin @{const_name "less_eq"} dummyT
- val prems_pairs = map (dest_less_eq o HOLogic.dest_Trueprop) (prems_of rule)
+ val prems_pairs = map (dest_less_eq o HOLogic.dest_Trueprop) (Thm.prems_of rule)
val equal_prems = filter op= prems_pairs
val _ =
if null equal_prems then ()
else error "The rule contains reflexive assumptions."
val concl_pairs = rule
- |> concl_of
+ |> Thm.concl_of
|> HOLogic.dest_Trueprop
|> dest_less_eq
|> apply2 (snd o strip_comb)
@@ -351,7 +351,7 @@
let
fun find_eq_rule thm ctxt =
let
- val concl_rhs = (hd o get_args 1 o HOLogic.dest_Trueprop o concl_of) thm;
+ val concl_rhs = (hd o get_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm;
val rules = Item_Net.retrieve (Transfer.get_relator_eq_item_net ctxt) concl_rhs;
in
find_first (fn thm => Pattern.matches (Proof_Context.theory_of ctxt) (concl_rhs,
@@ -372,7 +372,7 @@
let
val pol_mono_rule = introduce_polarities mono_rule
val mono_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o
- dest_Const o head_of o HOLogic.dest_Trueprop o concl_of) pol_mono_rule
+ dest_Const o head_of o HOLogic.dest_Trueprop o Thm.concl_of) pol_mono_rule
val _ = if Symtab.defined (get_relator_distr_data' ctxt) mono_ruleT_name
then error ("Monotocity rule for type " ^ quote mono_ruleT_name ^ " is already_defined.")
else ()
@@ -389,7 +389,7 @@
fun add_distr_rule update_entry distr_rule ctxt =
let
val distr_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o
- dest_Const o head_of o HOLogic.dest_Trueprop o concl_of) distr_rule
+ dest_Const o head_of o HOLogic.dest_Trueprop o Thm.concl_of) distr_rule
in
if Symtab.defined (get_relator_distr_data' ctxt) distr_ruleT_name then
Data.map (map_relator_distr_data (Symtab.map_entry distr_ruleT_name (update_entry distr_rule)))
@@ -441,8 +441,8 @@
local
fun sanity_check rule =
let
- val assms = map (perhaps (try HOLogic.dest_Trueprop)) (prems_of rule)
- val concl = (perhaps (try HOLogic.dest_Trueprop)) (concl_of rule);
+ val assms = map (perhaps (try HOLogic.dest_Trueprop)) (Thm.prems_of rule)
+ val concl = (perhaps (try HOLogic.dest_Trueprop)) (Thm.concl_of rule);
val (lhs, rhs) =
(case concl of
Const (@{const_name less_eq}, _) $ (lhs as Const (@{const_name relcompp},_) $ _ $ _) $ rhs =>
@@ -477,7 +477,7 @@
fun add_distr_rule distr_rule ctxt =
let
val _ = sanity_check distr_rule
- val concl = (perhaps (try HOLogic.dest_Trueprop)) (concl_of distr_rule)
+ val concl = (perhaps (try HOLogic.dest_Trueprop)) (Thm.concl_of distr_rule)
in
(case concl of
Const (@{const_name less_eq}, _) $ (Const (@{const_name relcompp},_) $ _ $ _) $ _ =>
--- a/src/HOL/Tools/Lifting/lifting_setup.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_setup.ML Wed Mar 04 19:53:18 2015 +0100
@@ -98,7 +98,7 @@
in
fun define_pcr_cr_eq lthy pcr_rel_def =
let
- val lhs = (term_of o Thm.lhs_of) pcr_rel_def
+ val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def
val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type o List.last o binder_types o fastype_of) lhs
val args = (snd o strip_comb) lhs
@@ -109,7 +109,7 @@
val (fresh_var, ctxt) = yield_singleton Variable.invent_types sort ctxt
val thy = Proof_Context.theory_of ctxt
in
- ((cterm_of thy var, cterm_of thy (HOLogic.eq_const (TFree fresh_var))), ctxt)
+ ((Thm.cterm_of thy var, Thm.cterm_of thy (HOLogic.eq_const (TFree fresh_var))), ctxt)
end
val orig_lthy = lthy
@@ -120,7 +120,7 @@
|> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv
(Transfer.bottom_rewr_conv (Transfer.get_relator_eq lthy))))
in
- case (term_of o Thm.rhs_of) pcr_cr_eq of
+ case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of
Const (@{const_name "relcompp"}, _) $ Const (@{const_name HOL.eq}, _) $ _ =>
let
val thm =
@@ -172,13 +172,13 @@
fun quot_thm_sanity_check ctxt quot_thm =
let
val _ =
- if (nprems_of quot_thm > 0) then
+ if (Thm.nprems_of quot_thm > 0) then
raise QUOT_ERROR [Pretty.block
[Pretty.str "The Quotient theorem has extra assumptions:",
Pretty.brk 1,
Display.pretty_thm ctxt quot_thm]]
else ()
- val _ = quot_thm |> concl_of |> HOLogic.dest_Trueprop |> dest_Quotient
+ val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient
handle TERM _ => raise QUOT_ERROR
[Pretty.block
[Pretty.str "The Quotient theorem is not of the right form:",
@@ -290,7 +290,7 @@
val thy = Proof_Context.theory_of ctxt
val orig_ctxt = ctxt
val (fixed_goal, ctxt) = yield_singleton (import_terms_exclude not_fix) goal ctxt
- val init_goal = Goal.init (cterm_of thy fixed_goal)
+ val init_goal = Goal.init (Thm.cterm_of thy fixed_goal)
in
(singleton (Variable.export ctxt orig_ctxt) o Goal.conclude) (the (SINGLE tac init_goal))
end
@@ -307,7 +307,7 @@
val thy = Proof_Context.theory_of ctxt
val orig_ctxt = ctxt
val (fixed_goal, ctxt) = yield_singleton (Variable.import_terms true) goal ctxt
- val init_goal = Goal.init (cterm_of thy fixed_goal)
+ val init_goal = Goal.init (Thm.cterm_of thy fixed_goal)
val rules = Transfer.get_transfer_raw ctxt
val rules = constraint :: OO_rules @ rules
val tac =
@@ -318,8 +318,9 @@
fun make_goal pcr_def constr =
let
- val pred_name = (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o prop_of) constr
- val arg = (fst o Logic.dest_equals o prop_of) pcr_def
+ val pred_name =
+ (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.prop_of) constr
+ val arg = (fst o Logic.dest_equals o Thm.prop_of) pcr_def
in
HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg)
end
@@ -337,8 +338,9 @@
in
fn thm =>
let
- val prems = map HOLogic.dest_Trueprop (prems_of thm)
- val thm_name = (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o concl_of) thm
+ val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm)
+ val thm_name =
+ (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm
val non_trivial_assms = filter_out is_trivial_assm prems
in
if null non_trivial_assms then ()
@@ -378,9 +380,9 @@
@{thm id_transfer}
|> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1)
|> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold))
- val var = Var (hd (Term.add_vars (prop_of id_transfer) []))
+ val var = Var (hd (Term.add_vars (Thm.prop_of id_transfer) []))
val thy = Proof_Context.theory_of lthy
- val inst = [(cterm_of thy var, cterm_of thy parametrized_relator)]
+ val inst = [(Thm.cterm_of thy var, Thm.cterm_of thy parametrized_relator)]
val id_par_thm = Drule.cterm_instantiate inst id_transfer
in
Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm
@@ -403,7 +405,9 @@
fun fold_Domainp_pcrel pcrel_def thm =
let
- val ct = thm |> cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg
+ val ct =
+ thm |> Thm.cprop_of |> Drule.strip_imp_concl
+ |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg
val pcrel_def = Thm.incr_indexes (#maxidx (Thm.rep_cterm ct) + 1) pcrel_def
val thm = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm
handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def])
@@ -413,7 +417,7 @@
fun reduce_Domainp ctxt rules thm =
let
- val goal = thm |> prems_of |> hd
+ val goal = thm |> Thm.prems_of |> hd
val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var
val reduced_assm =
reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
@@ -425,7 +429,7 @@
let
fun reduce_first_assm ctxt rules thm =
let
- val goal = thm |> prems_of |> hd
+ val goal = thm |> Thm.prems_of |> hd
val reduced_assm =
reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt
in
@@ -626,7 +630,7 @@
let
val transfer_attr = Attrib.internal (K Transfer.transfer_add)
val transfer_domain_attr = Attrib.internal (K Transfer.transfer_domain_add)
- val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o prop_of) typedef_thm
+ val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm
val (T_def, lthy) = define_crel rep_fun lthy
(**)
val T_def = Morphism.thm (Local_Theory.target_morphism lthy) T_def
@@ -742,12 +746,12 @@
fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy =
let
val input_thm = singleton (Attrib.eval_thms lthy) xthm
- val input_term = (HOLogic.dest_Trueprop o prop_of) input_thm
+ val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm
handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported."
fun sanity_check_reflp_thm reflp_thm =
let
- val reflp_tm = (HOLogic.dest_Trueprop o prop_of) reflp_thm
+ val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm
handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"."
in
case reflp_tm of
@@ -813,13 +817,13 @@
let
val pcrel_def = #pcrel_def pcr
val pcr_cr_eq = #pcr_cr_eq pcr
- val (def_lhs, _) = Logic.dest_equals (prop_of pcrel_def)
+ val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def)
handle TERM _ => raise PCR_ERROR [Pretty.block
[Pretty.str "The pcr definiton theorem is not a plain meta equation:",
Pretty.brk 1,
Display.pretty_thm ctxt pcrel_def]]
val pcr_const_def = head_of def_lhs
- val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of pcr_cr_eq))
+ val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq))
handle TERM _ => raise PCR_ERROR [Pretty.block
[Pretty.str "The pcr_cr equation theorem is not a plain equation:",
Pretty.brk 1,
@@ -939,7 +943,7 @@
val transfer_rel_name = transfer_rel |> dest_Const |> fst;
fun has_transfer_rel thm =
let
- val concl = thm |> concl_of |> HOLogic.dest_Trueprop
+ val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop
in
member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name
end
@@ -952,7 +956,7 @@
fun get_transfer_rel (qinfo : Lifting_Info.quotient) =
let
- fun get_pcrel pcr_def = pcr_def |> concl_of |> Logic.dest_equals |> fst |> head_of
+ fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of
in
if is_some (#pcr_info qinfo)
then get_pcrel (#pcrel_def (the (#pcr_info qinfo)))
--- a/src/HOL/Tools/Lifting/lifting_term.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_term.ML Wed Mar 04 19:53:18 2015 +0100
@@ -128,7 +128,7 @@
Pretty.str "found."]))
end
-fun is_id_quot thm = (prop_of thm = prop_of @{thm identity_quotient})
+fun is_id_quot thm = (Thm.prop_of thm = Thm.prop_of @{thm identity_quotient}) (* FIXME equality!? *)
fun zip_Tvars ctxt type_name rty_Tvars qty_Tvars =
case try (get_rel_quot_thm ctxt) type_name of
@@ -167,7 +167,7 @@
val rty = Type (type_name, rty_Tvars)
val qty = Type (type_name, qty_Tvars)
- val rel_quot_thm_concl = (Logic.strip_imp_concl o prop_of) rel_quot_thm
+ val rel_quot_thm_concl = (Logic.strip_imp_concl o Thm.prop_of) rel_quot_thm
val schematic_rel_absT = quot_term_absT rel_quot_thm_concl;
val thy = Proof_Context.theory_of ctxt
val absT = rty --> qty
@@ -180,7 +180,7 @@
val _ = Sign.typ_unify thy (schematic_rel_absT, schematic_absT) (Vartab.empty,maxidx)
handle Type.TUNIFY => equiv_univ_err ctxt schematic_rel_absT schematic_absT
val subs = raw_match (schematic_rel_absT, absT) Vartab.empty
- val rel_quot_thm_prems = (Logic.strip_imp_prems o prop_of) rel_quot_thm
+ val rel_quot_thm_prems = (Logic.strip_imp_prems o Thm.prop_of) rel_quot_thm
in
map (dest_funT o
Envir.subst_type subs o
@@ -278,7 +278,7 @@
val (_, qty_schematic) = quot_thm_rty_qty quot_thm
val match_env = Sign.typ_match thy (qty_schematic, qty) Vartab.empty
fun prep_ty thy (x, (S, ty)) =
- (ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
+ (Thm.ctyp_of thy (TVar (x, S)), Thm.ctyp_of thy ty)
val ty_inst = Vartab.fold (cons o (prep_ty thy)) match_env []
in
Thm.instantiate (ty_inst, []) quot_thm
@@ -372,7 +372,8 @@
then
let
val thy = Proof_Context.theory_of ctxt
- val instantiated_id_quot_thm = instantiate' [SOME (ctyp_of thy ty)] [] @{thm identity_quotient}
+ val instantiated_id_quot_thm =
+ instantiate' [SOME (Thm.ctyp_of thy ty)] [] @{thm identity_quotient}
in
(instantiated_id_quot_thm, (table, ctxt))
end
@@ -389,7 +390,7 @@
let
val thy = Proof_Context.theory_of ctxt
val (Q_t, ctxt') = get_fresh_Q_t ctxt
- val Q_thm = Thm.assume (cterm_of thy Q_t)
+ val Q_thm = Thm.assume (Thm.cterm_of thy Q_t)
val table' = (ty, Q_thm)::table
in
(Q_thm, (table', ctxt'))
@@ -422,7 +423,7 @@
(* Parametrization *)
local
- fun get_lhs rule = (Thm.dest_fun o Thm.dest_arg o strip_imp_concl o cprop_of) rule;
+ fun get_lhs rule = (Thm.dest_fun o Thm.dest_arg o strip_imp_concl o Thm.cprop_of) rule;
fun no_imp _ = raise CTERM ("no implication", []);
@@ -464,7 +465,7 @@
fun merge_transfer_relations ctxt ctm =
let
val ctm = Thm.dest_arg ctm
- val tm = term_of ctm
+ val tm = Thm.term_of ctm
val rel = (hd o get_args 2) tm
fun same_constants (Const (n1,_)) (Const (n2,_)) = n1 = n2
@@ -472,11 +473,11 @@
fun prove_extra_assms ctxt ctm distr_rule =
let
- fun prove_assm assm = try (Goal.prove ctxt [] [] (term_of assm))
+ fun prove_assm assm = try (Goal.prove ctxt [] [] (Thm.term_of assm))
(fn _ => SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt (Transfer.get_transfer_raw ctxt))) 1)
fun is_POS_or_NEG ctm =
- case (head_of o term_of o Thm.dest_arg) ctm of
+ case (head_of o Thm.term_of o Thm.dest_arg) ctm of
Const (@{const_name POS}, _) => true
| Const (@{const_name NEG}, _) => true
| _ => false
@@ -515,7 +516,7 @@
else
let
val pcrel_def = get_pcrel_def ctxt ((fst o dest_Type) qty)
- val pcrel_const = (head_of o fst o Logic.dest_equals o prop_of) pcrel_def
+ val pcrel_const = (head_of o fst o Logic.dest_equals o Thm.prop_of) pcrel_def
in
if same_constants pcrel_const (head_of trans_rel) then
let
@@ -547,7 +548,7 @@
let
fun parametrize_relation_conv ctm =
let
- val (rty, qty) = (relation_types o fastype_of) (term_of ctm)
+ val (rty, qty) = (relation_types o fastype_of) (Thm.term_of ctm)
in
if same_type_constrs (rty, qty) then
if forall op= (Targs rty ~~ Targs qty) then
--- a/src/HOL/Tools/Lifting/lifting_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Lifting/lifting_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -57,19 +57,19 @@
*)
fun quot_thm_rel quot_thm =
- case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
+ case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of
(rel, _, _, _) => rel
fun quot_thm_abs quot_thm =
- case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
+ case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of
(_, abs, _, _) => abs
fun quot_thm_rep quot_thm =
- case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
+ case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of
(_, _, rep, _) => rep
fun quot_thm_crel quot_thm =
- case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
+ case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of
(_, _, _, crel) => crel
fun quot_thm_rty_qty quot_thm =
@@ -87,7 +87,7 @@
Thm.implies_elim thm (Thm.assume assm)
end
-fun undisch_all thm = funpow (nprems_of thm) undisch thm
+fun undisch_all thm = funpow (Thm.nprems_of thm) undisch thm
fun is_fun_type (Type (@{type_name fun}, _)) = true
| is_fun_type _ = false
--- a/src/HOL/Tools/Meson/meson.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Meson/meson.ML Wed Mar 04 19:53:18 2015 +0100
@@ -101,13 +101,13 @@
fun first_order_resolve thA thB =
(case
try (fn () =>
- let val thy = theory_of_thm thA
- val tmA = concl_of thA
- val Const(@{const_name Pure.imp},_) $ tmB $ _ = prop_of thB
+ let val thy = Thm.theory_of_thm thA
+ val tmA = Thm.concl_of thA
+ val Const(@{const_name Pure.imp},_) $ tmB $ _ = Thm.prop_of thB
val tenv =
Pattern.first_order_match thy (tmB, tmA)
(Vartab.empty, Vartab.empty) |> snd
- val ct_pairs = map (apply2 (cterm_of thy) o term_pair_of) (Vartab.dest tenv)
+ val ct_pairs = map (apply2 (Thm.cterm_of thy) o term_pair_of) (Vartab.dest tenv)
in thA RS (cterm_instantiate ct_pairs thB) end) () of
SOME th => th
| NONE => raise THM ("first_order_resolve", 0, [thA, thB]))
@@ -152,10 +152,10 @@
(* Forward proof while preserving bound variables names *)
fun rename_bound_vars_RS th rl =
let
- val t = concl_of th
- val r = concl_of rl
+ val t = Thm.concl_of th
+ val r = Thm.concl_of rl
val th' = th RS Thm.rename_boundvars r (protect_bound_var_names r) rl
- val t' = concl_of th'
+ val t' = Thm.concl_of th'
in Thm.rename_boundvars t' (fix_bound_var_names t t') th' end
(*raises exception if no rules apply*)
@@ -171,11 +171,11 @@
problem, due to the (spurious) choices between projection and imitation. The
workaround is to instantiate "?P := (%c. ... c ... c ... c ...)" manually. *)
fun quant_resolve_tac ctxt th i st =
- case (concl_of st, prop_of th) of
+ case (Thm.concl_of st, Thm.prop_of th) of
(@{const Trueprop} $ (Var _ $ (c as Free _)), @{const Trueprop} $ _) =>
let
- val cc = cterm_of (theory_of_thm th) c
- val ct = Thm.dest_arg (cprop_of th)
+ val cc = Thm.cterm_of (Thm.theory_of_thm th) c
+ val ct = Thm.dest_arg (Thm.cprop_of th)
in resolve_tac ctxt [th] i (Drule.instantiate' [] [SOME (Thm.lambda cc ct)] st) end
| _ => resolve_tac ctxt [th] i st
@@ -228,7 +228,7 @@
| signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits)
| signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
-fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);
+fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (Thm.concl_of th)) ([],[]);
(*Literals like X=X are tautologous*)
fun taut_poslit (Const(@{const_name HOL.eq},_) $ t $ u) = t aconv u
@@ -258,7 +258,7 @@
fun refl_clause_aux 0 th = th
| refl_clause_aux n th =
- case HOLogic.dest_Trueprop (concl_of th) of
+ case HOLogic.dest_Trueprop (Thm.concl_of th) of
(Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) =>
refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*)
| (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name HOL.eq},_) $ t $ u)) $ _) =>
@@ -275,7 +275,7 @@
(*Simplify a clause by applying reflexivity to its negated equality literals*)
fun refl_clause th =
- let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th))
+ let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (Thm.concl_of th))
in zero_var_indexes (refl_clause_aux neqs th) end
handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*)
@@ -302,7 +302,7 @@
(*Remove duplicate literals, if there are any*)
fun nodups ctxt th =
- if has_duplicates (op =) (literals (prop_of th))
+ if has_duplicates (op =) (literals (Thm.prop_of th))
then nodups_aux ctxt [] th
else th;
@@ -353,7 +353,8 @@
fun freeze_spec th ctxt =
let
val cert = Proof_Context.cterm_of ctxt;
- val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (concl_of th))] ctxt;
+ val ([x], ctxt') =
+ Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (Thm.concl_of th))] ctxt;
val spec' = Thm.instantiate ([], [(spec_var, cert (Free (x, spec_varT)))]) spec;
in (th RS spec', ctxt') end
end;
@@ -371,10 +372,10 @@
fun cnf old_skolem_ths ctxt (th, ths) =
let val ctxt_ref = Unsynchronized.ref ctxt (* FIXME ??? *)
fun cnf_aux (th,ths) =
- if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
- else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th))
+ if not (can HOLogic.dest_Trueprop (Thm.prop_of th)) then ths (*meta-level: ignore*)
+ else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (Thm.prop_of th))
then nodups ctxt th :: ths (*no work to do, terminate*)
- else case head_of (HOLogic.dest_Trueprop (concl_of th)) of
+ else case head_of (HOLogic.dest_Trueprop (Thm.concl_of th)) of
Const (@{const_name HOL.conj}, _) => (*conjunction*)
cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
| Const (@{const_name All}, _) => (*universal quantifier*)
@@ -395,7 +396,7 @@
| _ => nodups ctxt th :: ths (*no work to do*)
and cnf_nil th = cnf_aux (th, [])
val cls =
- if has_too_many_clauses ctxt (concl_of th) then
+ if has_too_many_clauses ctxt (Thm.concl_of th) then
(trace_msg ctxt (fn () =>
"cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths)
else
@@ -417,7 +418,7 @@
(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
fun assoc_right th =
- if is_left (prop_of th) then assoc_right (th RS disj_assoc)
+ if is_left (Thm.prop_of th) then assoc_right (th RS disj_assoc)
else th;
(*Must check for negative literal first!*)
@@ -439,7 +440,7 @@
(*Create a meta-level Horn clause*)
fun make_horn crules th =
- if ok4horn (concl_of th)
+ if ok4horn (Thm.concl_of th)
then make_horn crules (tryres(th,crules)) handle THM _ => th
else th;
@@ -449,7 +450,7 @@
let fun rots (0,_) = hcs
| rots (k,th) = zero_var_indexes (make_horn crules th) ::
rots(k-1, assoc_right (th RS disj_comm))
- in case nliterals(prop_of th) of
+ in case nliterals(Thm.prop_of th) of
1 => th::hcs
| n => rots(n, assoc_right th)
end;
@@ -461,7 +462,7 @@
in fn ths => #2 (fold_rev name1 ths (length ths, [])) end;
(*Is the given disjunction an all-negative support clause?*)
-fun is_negative th = forall (not o #1) (literals (prop_of th));
+fun is_negative th = forall (not o #1) (literals (Thm.prop_of th));
val neg_clauses = filter is_negative;
@@ -489,12 +490,12 @@
TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st)
handle THM _ => TRYING_eq_assume_tac (i-1) st;
-fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st;
+fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (Thm.nprems_of st) st;
(*Loop checking: FAIL if trying to prove the same thing twice
-- if *ANY* subgoal has repeated literals*)
fun check_tac st =
- if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st)
+ if exists (fn prem => has_reps (rhyps(prem,[]))) (Thm.prems_of st)
then Seq.empty else Seq.single st;
@@ -506,7 +507,7 @@
(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)
fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz;
-fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0;
+fun size_of_subgoals st = fold_rev addconcl (Thm.prems_of st) 0;
(*Negation Normal Form*)
@@ -518,7 +519,7 @@
| ok4nnf _ = false;
fun make_nnf1 ctxt th =
- if ok4nnf (concl_of th)
+ if ok4nnf (Thm.concl_of th)
then make_nnf1 ctxt (tryres(th, nnf_rls))
handle THM ("tryres", _, _) =>
forward_res ctxt (make_nnf1 ctxt)
@@ -552,9 +553,10 @@
#> simplify (put_simpset nnf_ss ctxt)
#> rewrite_rule ctxt @{thms Let_def [abs_def]}
-fun make_nnf ctxt th = case prems_of th of
+fun make_nnf ctxt th =
+ (case Thm.prems_of th of
[] => th |> presimplify ctxt |> make_nnf1 ctxt
- | _ => raise THM ("make_nnf: premises in argument", 0, [th]);
+ | _ => raise THM ("make_nnf: premises in argument", 0, [th]));
fun choice_theorems thy =
try (Global_Theory.get_thm thy) "Hilbert_Choice.choice" |> the_list
@@ -564,7 +566,7 @@
fun skolemize_with_choice_theorems ctxt choice_ths =
let
fun aux th =
- if not (has_conns [@{const_name Ex}] (prop_of th)) then
+ if not (has_conns [@{const_name Ex}] (Thm.prop_of th)) then
th
else
tryres (th, choice_ths @
@@ -615,35 +617,35 @@
val ext_cong_neq = @{thm ext_cong_neq}
val F_ext_cong_neq =
- Term.add_vars (prop_of @{thm ext_cong_neq}) []
+ Term.add_vars (Thm.prop_of @{thm ext_cong_neq}) []
|> filter (fn ((s, _), _) => s = "F")
|> the_single |> Var
(* Strengthens "f g ~= f h" to "f g ~= f h & (EX x. g x ~= h x)". *)
fun cong_extensionalize_thm thy th =
- case concl_of th of
+ (case Thm.concl_of th of
@{const Trueprop} $ (@{const Not}
$ (Const (@{const_name HOL.eq}, Type (_, [T, _]))
$ (t as _ $ _) $ (u as _ $ _))) =>
(case get_F_pattern T t u of
SOME p =>
- let val inst = [apply2 (cterm_of thy) (F_ext_cong_neq, p)] in
+ let val inst = [apply2 (Thm.cterm_of thy) (F_ext_cong_neq, p)] in
th RS cterm_instantiate inst ext_cong_neq
end
| NONE => th)
- | _ => th
+ | _ => th)
(* Removes the lambdas from an equation of the form "t = (%x1 ... xn. u)". It
would be desirable to do this symmetrically but there's at least one existing
proof in "Tarski" that relies on the current behavior. *)
fun abs_extensionalize_conv ctxt ct =
- case term_of ct of
+ (case Thm.term_of ct of
Const (@{const_name HOL.eq}, _) $ _ $ Abs _ =>
ct |> (Conv.rewr_conv @{thm fun_eq_iff [THEN eq_reflection]}
then_conv abs_extensionalize_conv ctxt)
| _ $ _ => Conv.comb_conv (abs_extensionalize_conv ctxt) ct
| Abs _ => Conv.abs_conv (abs_extensionalize_conv o snd) ctxt ct
- | _ => Conv.all_conv ct
+ | _ => Conv.all_conv ct)
val abs_extensionalize_thm = Conv.fconv_rule o abs_extensionalize_conv
@@ -673,7 +675,7 @@
in Variable.export ctxt' ctxt cnfs @ cls end;
(*Sort clauses by number of literals*)
-fun fewerlits (th1, th2) = nliterals (prop_of th1) < nliterals (prop_of th2)
+fun fewerlits (th1, th2) = nliterals (Thm.prop_of th1) < nliterals (Thm.prop_of th2)
(*Make clauses from a list of theorems, previously Skolemized and put into nnf.
The resulting clauses are HOL disjunctions.*)
--- a/src/HOL/Tools/Meson/meson_clausify.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Meson/meson_clausify.ML Wed Mar 04 19:53:18 2015 +0100
@@ -35,20 +35,20 @@
(**** Transformation of Elimination Rules into First-Order Formulas****)
-val cfalse = cterm_of @{theory HOL} @{term False};
-val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop @{term False});
+val cfalse = Thm.cterm_of @{theory HOL} @{term False};
+val ctp_false = Thm.cterm_of @{theory HOL} (HOLogic.mk_Trueprop @{term False});
(* Converts an elim-rule into an equivalent theorem that does not have the
predicate variable. Leaves other theorems unchanged. We simply instantiate
the conclusion variable to False. (Cf. "transform_elim_prop" in
"Sledgehammer_Util".) *)
fun transform_elim_theorem th =
- case concl_of th of (*conclusion variable*)
- @{const Trueprop} $ (v as Var (_, @{typ bool})) =>
- Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
- | v as Var(_, @{typ prop}) =>
- Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
- | _ => th
+ (case Thm.concl_of th of (*conclusion variable*)
+ @{const Trueprop} $ (v as Var (_, @{typ bool})) =>
+ Thm.instantiate ([], [(Thm.cterm_of @{theory HOL} v, cfalse)]) th
+ | v as Var(_, @{typ prop}) =>
+ Thm.instantiate ([], [(Thm.cterm_of @{theory HOL} v, ctp_false)]) th
+ | _ => th)
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
@@ -83,7 +83,7 @@
| dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
| dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss
| dec_sko _ rhss = rhss
- in dec_sko (prop_of th) [] end;
+ in dec_sko (Thm.prop_of th) [] end;
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****)
@@ -94,20 +94,20 @@
| is_quasi_lambda_free (Abs _) = false
| is_quasi_lambda_free _ = true
-val [f_B,g_B] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_B}));
-val [g_C,f_C] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_C}));
-val [f_S,g_S] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_S}));
+val [f_B,g_B] = map (Thm.cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_B}));
+val [g_C,f_C] = map (Thm.cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_C}));
+val [f_S,g_S] = map (Thm.cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_S}));
(* FIXME: Requires more use of cterm constructors. *)
fun abstract ct =
let
- val thy = theory_of_cterm ct
- val Abs(x,_,body) = term_of ct
- val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct)
- val cxT = ctyp_of thy xT
- val cbodyT = ctyp_of thy bodyT
+ val thy = Thm.theory_of_cterm ct
+ val Abs(x,_,body) = Thm.term_of ct
+ val Type(@{type_name fun}, [xT,bodyT]) = Thm.typ_of (Thm.ctyp_of_term ct)
+ val cxT = Thm.ctyp_of thy xT
+ val cbodyT = Thm.ctyp_of thy bodyT
fun makeK () =
- instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)]
+ instantiate' [SOME cxT, SOME cbodyT] [SOME (Thm.cterm_of thy body)]
@{thm abs_K}
in
case body of
@@ -118,27 +118,28 @@
| rator$rand =>
if Term.is_dependent rator then (*C or S*)
if Term.is_dependent rand then (*S*)
- let val crator = cterm_of thy (Abs(x,xT,rator))
- val crand = cterm_of thy (Abs(x,xT,rand))
+ let val crator = Thm.cterm_of thy (Abs(x,xT,rator))
+ val crand = Thm.cterm_of thy (Abs(x,xT,rand))
val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
- val (_,rhs) = Thm.dest_equals (cprop_of abs_S')
+ val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S')
in
Thm.transitive abs_S' (Conv.binop_conv abstract rhs)
end
else (*C*)
- let val crator = cterm_of thy (Abs(x,xT,rator))
- val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
- val (_,rhs) = Thm.dest_equals (cprop_of abs_C')
+ let val crator = Thm.cterm_of thy (Abs(x,xT,rator))
+ val abs_C' =
+ cterm_instantiate [(f_C,crator),(g_C,Thm.cterm_of thy rand)] @{thm abs_C}
+ val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C')
in
Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)
end
else if Term.is_dependent rand then (*B or eta*)
if rand = Bound 0 then Thm.eta_conversion ct
else (*B*)
- let val crand = cterm_of thy (Abs(x,xT,rand))
- val crator = cterm_of thy rator
+ let val crand = Thm.cterm_of thy (Abs(x,xT,rand))
+ val crator = Thm.cterm_of thy rator
val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
- val (_,rhs) = Thm.dest_equals (cprop_of abs_B')
+ val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B')
in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end
else makeK()
| _ => raise Fail "abstract: Bad term"
@@ -146,13 +147,13 @@
(* Traverse a theorem, remplacing lambda-abstractions with combinators. *)
fun introduce_combinators_in_cterm ct =
- if is_quasi_lambda_free (term_of ct) then
+ if is_quasi_lambda_free (Thm.term_of ct) then
Thm.reflexive ct
- else case term_of ct of
+ else case Thm.term_of ct of
Abs _ =>
let
val (cv, cta) = Thm.dest_abs NONE ct
- val (v, _) = dest_Free (term_of cv)
+ val (v, _) = dest_Free (Thm.term_of cv)
val u_th = introduce_combinators_in_cterm cta
val cu = Thm.rhs_of u_th
val comb_eq = abstract (Thm.lambda cv cu)
@@ -164,12 +165,12 @@
end
fun introduce_combinators_in_theorem ctxt th =
- if is_quasi_lambda_free (prop_of th) then
+ if is_quasi_lambda_free (Thm.prop_of th) then
th
else
let
val th = Drule.eta_contraction_rule th
- val eqth = introduce_combinators_in_cterm (cprop_of th)
+ val eqth = introduce_combinators_in_cterm (Thm.cprop_of th)
in Thm.equal_elim eqth th end
handle THM (msg, _, _) =>
(warning ("Error in the combinator translation of " ^ Display.string_of_thm ctxt th ^
@@ -178,7 +179,7 @@
TrueI)
(*cterms are used throughout for efficiency*)
-val cTrueprop = cterm_of @{theory HOL} HOLogic.Trueprop;
+val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
(*Given an abstraction over n variables, replace the bound variables by free
ones. Return the body, along with the list of free variables.*)
@@ -193,15 +194,15 @@
fun old_skolem_theorem_of_def ctxt rhs0 =
let
val thy = Proof_Context.theory_of ctxt
- val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> cterm_of thy
+ val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of thy
val rhs' = rhs |> Thm.dest_comb |> snd
val (ch, frees) = c_variant_abs_multi (rhs', [])
- val (hilbert, cabs) = ch |> Thm.dest_comb |>> term_of
+ val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of
val T =
case hilbert of
Const (_, Type (@{type_name fun}, [_, T])) => T
| _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert])
- val cex = cterm_of thy (HOLogic.exists_const T)
+ val cex = Thm.cterm_of thy (HOLogic.exists_const T)
val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs)
val conc =
Drule.list_comb (rhs, frees)
@@ -219,7 +220,7 @@
fun to_definitional_cnf_with_quantifiers ctxt th =
let
- val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (prop_of th))
+ val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (Thm.prop_of th))
val eqth = eqth RS @{thm eq_reflection}
val eqth = eqth RS @{thm TruepropI}
in Thm.equal_elim eqth th end
@@ -271,7 +272,7 @@
fun zap pos ct =
ct
- |> (case term_of ct of
+ |> (case Thm.term_of ct of
Const (s, _) $ Abs (s', _, _) =>
if s = @{const_name Pure.all} orelse s = @{const_name All} orelse
s = @{const_name Ex} then
@@ -325,23 +326,23 @@
skolemize choice_ths
val discharger_th = th |> pull_out
val discharger_th =
- discharger_th |> has_too_many_clauses ctxt (concl_of discharger_th)
+ discharger_th |> has_too_many_clauses ctxt (Thm.concl_of discharger_th)
? (to_definitional_cnf_with_quantifiers ctxt
#> pull_out)
val zapped_th =
- discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no
+ discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no
|> (if no_choice then
- Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of
+ Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of
else
- cterm_of thy)
+ Thm.cterm_of thy)
|> zap true
val fixes =
- [] |> Term.add_free_names (prop_of zapped_th)
+ [] |> Term.add_free_names (Thm.prop_of zapped_th)
|> filter is_zapped_var_name
val ctxt' = ctxt |> Variable.add_fixes_direct fixes
val fully_skolemized_t =
zapped_th |> singleton (Variable.export ctxt' ctxt)
- |> cprop_of |> Thm.dest_equals |> snd |> term_of
+ |> Thm.cprop_of |> Thm.dest_equals |> snd |> Thm.term_of
in
if exists_subterm (fn Var ((s, _), _) =>
String.isPrefix new_skolem_var_prefix s
@@ -349,7 +350,7 @@
let
val (fully_skolemized_ct, ctxt) =
Variable.import_terms true [fully_skolemized_t] ctxt
- |>> the_single |>> cterm_of thy
+ |>> the_single |>> Thm.cterm_of thy
in
(SOME (discharger_th, fully_skolemized_ct),
(Thm.assume fully_skolemized_ct, ctxt))
@@ -358,7 +359,7 @@
(NONE, (th, ctxt))
end
else
- (NONE, (th |> has_too_many_clauses ctxt (concl_of th)
+ (NONE, (th |> has_too_many_clauses ctxt (Thm.concl_of th)
? to_definitional_cnf_with_quantifiers ctxt, ctxt))
end
@@ -376,14 +377,14 @@
th ctxt
val (cnf_ths, ctxt) = clausify nnf_th
fun intr_imp ct th =
- Thm.instantiate ([], map (apply2 (cterm_of thy))
+ Thm.instantiate ([], map (apply2 (Thm.cterm_of thy))
[(Var (("i", 0), @{typ nat}),
HOLogic.mk_nat ax_no)])
(zero_var_indexes @{thm skolem_COMBK_D})
RS Thm.implies_intr ct th
in
(opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0)
- ##> (term_of #> HOLogic.dest_Trueprop
+ ##> (Thm.term_of #> HOLogic.dest_Trueprop
#> singleton (Variable.export_terms ctxt ctxt0))),
cnf_ths |> map (combinators ? introduce_combinators_in_theorem ctxt
#> (case opt of SOME (_, ct) => intr_imp ct | NONE => I))
--- a/src/HOL/Tools/Metis/metis_generate.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Metis/metis_generate.ML Wed Mar 04 19:53:18 2015 +0100
@@ -216,7 +216,7 @@
(0 upto length fact_clauses - 1) fact_clauses
val (old_skolems, props) =
fold_rev (fn (name, th) => fn (old_skolems, props) =>
- th |> prop_of |> Logic.strip_imp_concl
+ th |> Thm.prop_of |> Logic.strip_imp_concl
|> conceal_old_skolem_terms (length clauses) old_skolems
||> (lam_trans = liftingN orelse lam_trans = lam_liftingN) ? eliminate_lam_wrappers
||> (fn prop => (name, prop) :: props))
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -88,7 +88,7 @@
the (AList.lookup (uncurry Metis_Thm.equal) th_pairs fth)
handle Option.Option => raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth)
-fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx))
+fun cterm_incr_types thy idx = Thm.cterm_of thy o map_types (Logic.incr_tvar idx)
(* INFERENCE RULE: AXIOM *)
@@ -103,8 +103,8 @@
fun inst_excluded_middle thy i_atom =
let
val th = EXCLUDED_MIDDLE
- val [vx] = Term.add_vars (prop_of th) []
- val substs = [(cterm_of thy (Var vx), cterm_of thy i_atom)]
+ val [vx] = Term.add_vars (Thm.prop_of th) []
+ val substs = [(Thm.cterm_of thy (Var vx), Thm.cterm_of thy i_atom)]
in
cterm_instantiate substs th
end
@@ -122,7 +122,7 @@
let
val thy = Proof_Context.theory_of ctxt
val i_th = lookth th_pairs th
- val i_th_vars = Term.add_vars (prop_of i_th) []
+ val i_th_vars = Term.add_vars (Thm.prop_of i_th) []
fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
fun subst_translation (x,y) =
@@ -131,7 +131,7 @@
(* We call "polish_hol_terms" below. *)
val t = hol_term_of_metis ctxt type_enc sym_tab y
in
- SOME (cterm_of thy (Var v), t)
+ SOME (Thm.cterm_of thy (Var v), t)
end
handle Option.Option =>
(trace_msg ctxt (fn () =>
@@ -160,8 +160,8 @@
val _ = trace_msg ctxt (fn () =>
cat_lines ("subst_translations:" ::
(substs' |> map (fn (x, y) =>
- Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
- Syntax.string_of_term ctxt (term_of y)))))
+ Syntax.string_of_term ctxt (Thm.term_of x) ^ " |-> " ^
+ Syntax.string_of_term ctxt (Thm.term_of y)))))
in
cterm_instantiate substs' i_th
end
@@ -175,7 +175,7 @@
let
val tvs = Term.add_tvars (Thm.full_prop_of th) []
val thy = Thm.theory_of_thm th
- fun inc_tvar ((a, i), s) = apply2 (ctyp_of thy) (TVar ((a, i), s), TVar ((a, i + inc), s))
+ fun inc_tvar ((a, i), s) = apply2 (Thm.ctyp_of thy) (TVar ((a, i), s), TVar ((a, i + inc), s))
in
Thm.instantiate (map inc_tvar tvs, []) th
end
@@ -188,7 +188,7 @@
val tha = incr_type_indexes (1 + Thm.maxidx_of thb) tha
fun res (tha, thb) =
(case Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true}
- (false, tha, nprems_of tha) i thb
+ (false, tha, Thm.nprems_of tha) i thb
|> Seq.list_of |> distinct Thm.eq_thm of
[th] => th
| _ =>
@@ -207,13 +207,13 @@
let
val thy = Proof_Context.theory_of ctxt
val ps = []
- |> fold (Term.add_vars o prop_of) [tha, thb]
+ |> fold (Term.add_vars o Thm.prop_of) [tha, thb]
|> AList.group (op =)
|> maps (fn ((s, _), T :: Ts) => map (fn T' => (Free (s, T), Free (s, T'))) Ts)
|> rpair (Envir.empty ~1)
|-> fold (Pattern.unify (Context.Proof ctxt))
|> Envir.type_env |> Vartab.dest
- |> map (fn (x, (S, T)) => apply2 (ctyp_of thy) (TVar (x, S), T))
+ |> map (fn (x, (S, T)) => apply2 (Thm.ctyp_of thy) (TVar (x, S), T))
in
(* The unifier, which is invoked from "Thm.bicompose", will sometimes refuse to unify
"?a::?'a" with "?a::?'b" or "?a::nat" and throw a "TERM" exception (with "add_ffpair" as
@@ -248,7 +248,7 @@
(* Permute a rule's premises to move the i-th premise to the last position. *)
fun make_last i th =
- let val n = nprems_of th in
+ let val n = Thm.nprems_of th in
if i >= 1 andalso i <= n then Thm.permute_prems (i - 1) 1 th
else raise THM ("select_literal", i, [th])
end
@@ -259,7 +259,7 @@
don't use this trick in general because it makes the proof object uglier than
necessary. FIXME. *)
fun negate_head ctxt th =
- if exists (fn t => t aconv @{prop "~ False"}) (prems_of th) then
+ if exists (fn t => t aconv @{prop "~ False"}) (Thm.prems_of th) then
(th RS @{thm select_FalseI})
|> fold (rewrite_rule ctxt o single) @{thms not_atomize_select atomize_not_select}
else
@@ -286,11 +286,11 @@
singleton (hol_terms_of_metis ctxt type_enc concealed sym_tab) (Metis_Term.Fn atom)
val _ = trace_msg ctxt (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atom)
in
- (case index_of_literal (s_not i_atom) (prems_of i_th1) of
+ (case index_of_literal (s_not i_atom) (Thm.prems_of i_th1) of
0 => (trace_msg ctxt (fn () => "Failed to find literal in \"th1\""); i_th1)
| j1 =>
(trace_msg ctxt (fn () => " index th1: " ^ string_of_int j1);
- (case index_of_literal i_atom (prems_of i_th2) of
+ (case index_of_literal i_atom (Thm.prems_of i_th2) of
0 => (trace_msg ctxt (fn () => "Failed to find literal in \"th2\""); i_th2)
| j2 =>
(trace_msg ctxt (fn () => " index th2: " ^ string_of_int j2);
@@ -303,7 +303,7 @@
val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
-val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
+val refl_x = Thm.cterm_of @{theory} (Var (hd (Term.add_vars (Thm.prop_of REFL_THM) [])));
val refl_idx = 1 + Thm.maxidx_of REFL_THM;
fun refl_inference ctxt type_enc concealed sym_tab t =
@@ -374,8 +374,8 @@
val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill-typed but gives right max*)
val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
val _ = trace_msg ctxt (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
- val eq_terms = map (apply2 (cterm_of thy))
- (ListPair.zip (Misc_Legacy.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+ val eq_terms = map (apply2 (Thm.cterm_of thy))
+ (ListPair.zip (Misc_Legacy.term_vars (Thm.prop_of subst'), [tm_abs, tm_subst, i_tm]))
in
cterm_instantiate eq_terms subst'
end
@@ -399,9 +399,9 @@
[] => th
| pairs =>
let
- val thy = theory_of_thm th
- val cert = cterm_of thy
- val certT = ctyp_of thy
+ val thy = Thm.theory_of_thm th
+ val cert = Thm.cterm_of thy
+ val certT = Thm.ctyp_of thy
val (tyenv, tenv) = fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty)
fun mkT (v, (S, T)) = (certT (TVar (v, S)), certT T)
@@ -429,13 +429,13 @@
let
val num_metis_lits =
count is_metis_literal_genuine (Metis_LiteralSet.toList (Metis_Thm.clause fol_th))
- val num_isabelle_lits = count is_isabelle_literal_genuine (prems_of th)
+ val num_isabelle_lits = count is_isabelle_literal_genuine (Thm.prems_of th)
in
if num_metis_lits >= num_isabelle_lits then
th
else
let
- val (prems0, concl) = th |> prop_of |> Logic.strip_horn
+ val (prems0, concl) = th |> Thm.prop_of |> Logic.strip_horn
val prems = prems0 |> map normalize_literal |> distinct Term.aconv_untyped
val goal = Logic.list_implies (prems, concl)
val ctxt' = fold Thm.declare_hyps (#hyps (Thm.crep_thm th)) ctxt
@@ -454,7 +454,7 @@
fun replay_one_inference ctxt type_enc concealed sym_tab (fol_th, inf)
th_pairs =
- if not (null th_pairs) andalso prop_of (snd (hd th_pairs)) aconv @{prop False} then
+ if not (null th_pairs) andalso Thm.prop_of (snd (hd th_pairs)) aconv @{prop False} then
(* Isabelle sometimes identifies literals (premises) that are distinct in
Metis (e.g., because of type variables). We give the Isabelle proof the
benefice of the doubt. *)
@@ -481,7 +481,7 @@
where the nonvariables are goal parameters. *)
fun unify_first_prem_with_concl thy i th =
let
- val goal = Logic.get_goal (prop_of th) i |> Envir.beta_eta_contract
+ val goal = Logic.get_goal (Thm.prop_of th) i |> Envir.beta_eta_contract
val prem = goal |> Logic.strip_assums_hyp |> hd
val concl = goal |> Logic.strip_assums_concl
@@ -522,7 +522,7 @@
| _ => I)
val t_inst =
- [] |> try (unify_terms (prem, concl) #> map (apply2 (cterm_of thy)))
+ [] |> try (unify_terms (prem, concl) #> map (apply2 (Thm.cterm_of thy)))
|> the_default [] (* FIXME *)
in
cterm_instantiate t_inst th
@@ -543,7 +543,7 @@
let
val thy = Proof_Context.theory_of ctxt
- val params = Logic.strip_params (Logic.get_goal (prop_of st) i) |> rev
+ val params = Logic.strip_params (Logic.get_goal (Thm.prop_of st) i) |> rev
fun repair (t as (Var ((s, _), _))) =
(case find_index (fn (s', _) => s' = s) params of
@@ -561,7 +561,7 @@
val t' = t |> repair |> fold (absdummy o snd) params
fun do_instantiate th =
- (case Term.add_vars (prop_of th) []
+ (case Term.add_vars (Thm.prop_of th) []
|> filter_out ((Meson_Clausify.is_zapped_var_name orf is_metis_fresh_variable) o fst
o fst) of
[] => th
@@ -576,8 +576,9 @@
Envir.Envir {maxidx = Vartab.fold (Integer.max o snd o fst) tyenv 0,
tenv = Vartab.empty, tyenv = tyenv}
val ty_inst =
- Vartab.fold (fn (x, (S, T)) => cons (apply2 (ctyp_of thy) (TVar (x, S), T))) tyenv []
- val t_inst = [apply2 (cterm_of thy o Envir.norm_term env) (Var var, t')]
+ Vartab.fold (fn (x, (S, T)) => cons (apply2 (Thm.ctyp_of thy) (TVar (x, S), T)))
+ tyenv []
+ val t_inst = [apply2 (Thm.cterm_of thy o Envir.norm_term env) (Var var, t')]
in
Drule.instantiate_normalize (ty_inst, t_inst) th
end
@@ -639,7 +640,7 @@
specified axioms. The axioms have leading "All" and "Ex" quantifiers, which
must be eliminated first. *)
fun discharge_skolem_premises ctxt axioms prems_imp_false =
- if prop_of prems_imp_false aconv @{prop False} then
+ if Thm.prop_of prems_imp_false aconv @{prop False} then
prems_imp_false
else
let
@@ -685,7 +686,7 @@
clusters_in_term true t |> cluster_no > 1 ? cons (ax_no, cluster_no - 1))
| _ => raise TERM ("discharge_skolem_premises: Expected Var", [var]))
- val prems = Logic.strip_imp_prems (prop_of prems_imp_false)
+ val prems = Logic.strip_imp_prems (Thm.prop_of prems_imp_false)
val substs = prems |> map2 subst_info_of_prem (1 upto length prems)
|> sort (int_ord o apply2 fst)
val depss = maps (map_filter deps_of_term_subst o snd o snd o snd) substs
--- a/src/HOL/Tools/Metis/metis_tactic.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Metis/metis_tactic.ML Wed Mar 04 19:53:18 2015 +0100
@@ -33,7 +33,7 @@
(* Designed to work also with monomorphic instances of polymorphic theorems. *)
fun have_common_thm ths1 ths2 =
- exists (member (Term.aconv_untyped o apply2 prop_of) ths1) (map Meson.make_meta_clause ths2)
+ exists (member (Term.aconv_untyped o apply2 Thm.prop_of) ths1) (map Meson.make_meta_clause ths2)
(*Determining which axiom clauses are actually used*)
fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th)
@@ -48,12 +48,12 @@
(case hol_clause_of_metis ctxt type_enc sym_tab concealed mth of
Const (@{const_name HOL.eq}, _) $ _ $ t =>
let
- val ct = cterm_of thy t
- val cT = ctyp_of_term ct
+ val ct = Thm.cterm_of thy t
+ val cT = Thm.ctyp_of_term ct
in refl |> Drule.instantiate' [SOME cT] [SOME ct] end
| Const (@{const_name disj}, _) $ t1 $ t2 =>
(if can HOLogic.dest_not t1 then t2 else t1)
- |> HOLogic.mk_Trueprop |> cterm_of thy |> Thm.trivial
+ |> HOLogic.mk_Trueprop |> Thm.cterm_of thy |> Thm.trivial
| _ => raise Fail "expected reflexive or trivial clause")
end
|> Meson.make_meta_clause
@@ -63,7 +63,7 @@
val thy = Proof_Context.theory_of ctxt
val tac = rewrite_goals_tac ctxt @{thms lambda_def [abs_def]} THEN resolve_tac ctxt [refl] 1
val t = hol_clause_of_metis ctxt type_enc sym_tab concealed mth
- val ct = cterm_of thy (HOLogic.mk_Trueprop t)
+ val ct = Thm.cterm_of thy (HOLogic.mk_Trueprop t)
in Goal.prove_internal ctxt [] ct (K tac) |> Meson.make_meta_clause end
fun add_vars_and_frees (t $ u) = fold (add_vars_and_frees) [t, u]
@@ -73,16 +73,16 @@
| add_vars_and_frees _ = I
fun introduce_lam_wrappers ctxt th =
- if Meson_Clausify.is_quasi_lambda_free (prop_of th) then
+ if Meson_Clausify.is_quasi_lambda_free (Thm.prop_of th) then
th
else
let
val thy = Proof_Context.theory_of ctxt
fun conv first ctxt ct =
- if Meson_Clausify.is_quasi_lambda_free (term_of ct) then
+ if Meson_Clausify.is_quasi_lambda_free (Thm.term_of ct) then
Thm.reflexive ct
else
- (case term_of ct of
+ (case Thm.term_of ct of
Abs (_, _, u) =>
if first then
(case add_vars_and_frees u [] of
@@ -90,18 +90,18 @@
Conv.abs_conv (conv false o snd) ctxt ct
|> (fn th => Meson.first_order_resolve th @{thm Metis.eq_lambdaI})
| v :: _ =>
- Abs (Name.uu, fastype_of v, abstract_over (v, term_of ct)) $ v
- |> cterm_of thy
+ Abs (Name.uu, fastype_of v, abstract_over (v, Thm.term_of ct)) $ v
+ |> Thm.cterm_of thy
|> Conv.comb_conv (conv true ctxt))
else
Conv.abs_conv (conv false o snd) ctxt ct
| Const (@{const_name Meson.skolem}, _) $ _ => Thm.reflexive ct
| _ => Conv.comb_conv (conv true ctxt) ct)
- val eq_th = conv true ctxt (cprop_of th)
+ val eq_th = conv true ctxt (Thm.cprop_of th)
(* We replace the equation's left-hand side with a beta-equivalent term
so that "Thm.equal_elim" works below. *)
- val t0 $ _ $ t2 = prop_of eq_th
- val eq_ct = t0 $ prop_of th $ t2 |> cterm_of thy
+ val t0 $ _ $ t2 = Thm.prop_of eq_th
+ val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.cterm_of thy
val eq_th' = Goal.prove_internal ctxt [] eq_ct (K (resolve_tac ctxt [eq_th] 1))
in Thm.equal_elim eq_th' th end
@@ -189,7 +189,7 @@
("Falling back on " ^ quote (metis_call (hd fallback_type_encs) lam_trans) ^ "...");
FOL_SOLVE unused fallback_type_encs lam_trans ctxt cls ths0)
in
- (case filter (fn t => prop_of t aconv @{prop False}) cls of
+ (case filter (fn t => Thm.prop_of t aconv @{prop False}) cls of
false_th :: _ => [false_th RS @{thm FalseE}]
| [] =>
(case Metis_Resolution.loop (Metis_Resolution.new (resolution_params ordering)
@@ -197,7 +197,7 @@
Metis_Resolution.Contradiction mth =>
let
val _ = trace_msg ctxt (fn () => "METIS RECONSTRUCTION START: " ^ Metis_Thm.toString mth)
- val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
+ val ctxt' = fold Variable.declare_constraints (map Thm.prop_of cls) ctxt
(*add constraints arising from converting goal to clause form*)
val proof = Metis_Proof.proof mth
val result = fold (replay_one_inference ctxt' type_enc concealed sym_tab) proof axioms
@@ -243,7 +243,7 @@
fun preskolem_tac ctxt st0 =
(if exists (Meson.has_too_many_clauses ctxt)
- (Logic.prems_of_goal (prop_of st0) 1) then
+ (Logic.prems_of_goal (Thm.prop_of st0) 1) then
Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex} ctxt) 1
THEN CNF.cnfx_rewrite_tac ctxt 1
else
@@ -269,7 +269,7 @@
prevent "Subgoal.FOCUS" from freezing the type variables. We don't do it for nonschematic facts
"X" because this breaks a few proofs (in the rare and subtle case where a proof relied on
extensionality not being applied) and brings few benefits. *)
-val has_tvar = exists_type (exists_subtype (fn TVar _ => true | _ => false)) o prop_of
+val has_tvar = exists_type (exists_subtype (fn TVar _ => true | _ => false)) o Thm.prop_of
fun metis_method ((override_type_encs, lam_trans), ths) ctxt facts =
let val (schem_facts, nonschem_facts) = List.partition has_tvar facts in
--- a/src/HOL/Tools/Nitpick/nitpick.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1065,14 +1065,14 @@
fun pick_nits_in_subgoal state params mode i step =
let
val ctxt = Proof.context_of state
- val t = state |> Proof.raw_goal |> #goal |> prop_of
+ val t = state |> Proof.raw_goal |> #goal |> Thm.prop_of
in
case Logic.count_prems t of
0 => (writeln "No subgoal!"; (noneN, []))
| n =>
let
val t = Logic.goal_params t i |> fst
- val assms = map term_of (Assumption.all_assms_of ctxt)
+ val assms = map Thm.term_of (Assumption.all_assms_of ctxt)
val (subst, assms, t) = extract_fixed_frees ctxt (assms, t)
in pick_nits_in_term state params mode i n step subst [] assms t end
end
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Mar 04 19:53:18 2015 +0100
@@ -598,7 +598,7 @@
{Rep, Abs_inverse, Rep_inverse, ...}) :: _ =>
SOME {abs_type = Logic.varifyT_global abs_type,
rep_type = Logic.varifyT_global rep_type, Abs_name = Abs_name,
- Rep_name = Rep_name, prop_of_Rep = prop_of Rep,
+ Rep_name = Rep_name, prop_of_Rep = Thm.prop_of Rep,
Abs_inverse = SOME Abs_inverse, Rep_inverse = SOME Rep_inverse}
| _ => NONE
@@ -805,7 +805,7 @@
val {qtyp, equiv_rel, equiv_thm, ...} =
the (Quotient_Info.lookup_quotients thy s)
val partial =
- case prop_of equiv_thm of
+ case Thm.prop_of equiv_thm of
@{const Trueprop} $ (Const (@{const_name equivp}, _) $ _) => false
| @{const Trueprop} $ (Const (@{const_name part_equivp}, _) $ _) => true
| _ => raise NOT_SUPPORTED "Ill-formed quotient type equivalence \
@@ -1349,7 +1349,7 @@
in
Theory.nodes_of thy
|> maps Thm.axioms_of
- |> map (apsnd (subst_atomic subst o prop_of))
+ |> map (apsnd (subst_atomic subst o Thm.prop_of))
|> sort (fast_string_ord o apply2 fst)
|> Ord_List.inter (fast_string_ord o apsnd fst) def_names
|> map snd
@@ -1364,8 +1364,8 @@
ctxt |> Spec_Rules.get
|> filter (curry (op =) Spec_Rules.Unknown o fst)
|> maps (snd o snd)
- |> filter_out (is_built_in_theory o theory_of_thm)
- |> map (subst_atomic subst o prop_of)
+ |> filter_out (is_built_in_theory o Thm.theory_of_thm)
+ |> map (subst_atomic subst o Thm.prop_of)
fun arity_of_built_in_const (s, T) =
if s = @{const_name If} then
@@ -1928,7 +1928,7 @@
fun const_def_tables ctxt subst ts =
(def_table_for
- (map prop_of (rev (Named_Theorems.get ctxt @{named_theorems nitpick_unfold}))) subst,
+ (map Thm.prop_of (rev (Named_Theorems.get ctxt @{named_theorems nitpick_unfold}))) subst,
fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
(map pair_for_prop ts) Symtab.empty)
@@ -1938,22 +1938,22 @@
fold (append o paired_with_consts) ts [] |> AList.group (op =) |> Symtab.make
fun const_simp_table ctxt =
- def_table_for (map_filter (equationalize_term ctxt "nitpick_simp" o prop_of)
+ def_table_for (map_filter (equationalize_term ctxt "nitpick_simp" o Thm.prop_of)
(rev (Named_Theorems.get ctxt @{named_theorems nitpick_simp})))
fun const_psimp_table ctxt =
- def_table_for (map_filter (equationalize_term ctxt "nitpick_psimp" o prop_of)
+ def_table_for (map_filter (equationalize_term ctxt "nitpick_psimp" o Thm.prop_of)
(rev (Named_Theorems.get ctxt @{named_theorems nitpick_psimp})))
fun const_choice_spec_table ctxt subst =
- map (subst_atomic subst o prop_of)
+ map (subst_atomic subst o Thm.prop_of)
(rev (Named_Theorems.get ctxt @{named_theorems nitpick_choice_spec}))
|> const_nondef_table
fun inductive_intro_table ctxt subst def_tables =
let val thy = Proof_Context.theory_of ctxt in
def_table_for
- (maps (map (unfold_mutually_inductive_preds thy def_tables o prop_of)
+ (maps (map (unfold_mutually_inductive_preds thy def_tables o Thm.prop_of)
o snd o snd)
(filter (fn (cat, _) => cat = Spec_Rules.Inductive orelse
cat = Spec_Rules.Co_Inductive) (Spec_Rules.get ctxt))) subst
@@ -1962,7 +1962,7 @@
fun ground_theorem_table thy =
fold ((fn @{const Trueprop} $ t1 =>
is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1)
- | _ => I) o prop_of o snd) (Global_Theory.all_thms_of thy true) Inttab.empty
+ | _ => I) o Thm.prop_of o snd) (Global_Theory.all_thms_of thy true) Inttab.empty
fun ersatz_table ctxt =
#ersatz_table (Data.get (Context.Proof ctxt))
@@ -1979,7 +1979,7 @@
in
typedef_info ctxt (fst (dest_Type abs_T)) |> the
|> pairf #Abs_inverse #Rep_inverse
- |> apply2 (specialize_type thy x o prop_of o the)
+ |> apply2 (specialize_type thy x o Thm.prop_of o the)
||> single |> op ::
end
@@ -2149,7 +2149,7 @@
SOME wf => wf
| NONE =>
let
- val goal = prop |> cterm_of thy |> Goal.init
+ val goal = prop |> Thm.cterm_of thy |> Goal.init
val wf = exists (terminates_by ctxt tac_timeout goal)
termination_tacs
in Synchronized.change cached_wf_props (cons (prop, wf)); wf end
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1062,7 +1062,7 @@
|> writeln
else
()
- val goal = prop |> cterm_of thy |> Goal.init
+ val goal = prop |> Thm.cterm_of thy |> Goal.init
in
(goal |> SINGLE (DETERM_TIMEOUT auto_timeout (auto_tac ctxt))
|> the |> Goal.finish ctxt; true)
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1025,7 +1025,7 @@
let
val supers = Sign.complete_sort thy S
val class_axioms =
- maps (fn class => map prop_of (Axclass.get_info thy class |> #axioms
+ maps (fn class => map Thm.prop_of (Axclass.get_info thy class |> #axioms
handle ERROR _ => [])) supers
val monomorphic_class_axioms =
map (fn t => case Term.add_tvars t [] of
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -305,7 +305,7 @@
let
val (state, i, message) = f ()
val ctxt = Proof.context_of state
- val goal = Logic.get_goal (prop_of (#goal (Proof.goal state))) i
+ val goal = Logic.get_goal (Thm.prop_of (#goal (Proof.goal state))) i
val hash = String.substring (SHA1.rep (SHA1.digest (hackish_string_of_term ctxt goal)), 0, 12)
in
File.append (Path.explode "$ISABELLE_HOME_USER/spy_nitpick")
--- a/src/HOL/Tools/Old_Datatype/old_datatype.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML Wed Mar 04 19:53:18 2015 +0100
@@ -27,7 +27,7 @@
(** auxiliary **)
val distinct_lemma = @{lemma "f x \<noteq> f y ==> x \<noteq> y" by iprover};
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (Thm.prems_of distinct_lemma);
fun exh_thm_of (dt_info : Old_Datatype_Aux.info Symtab.table) tname =
#exhaust (the (Symtab.lookup dt_info tname));
@@ -246,10 +246,10 @@
fun dt_constr_defs (((((_, (_, _, constrs)), tname), typedef: Typedef.info), T), constr_syntax)
(thy, defs, eqns, rep_congs, dist_lemmas) =
let
- val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
- val rep_const = cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
- val cong' = cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong;
- val dist = cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma;
+ val _ $ (_ $ (cong_f $ _) $ _) = Thm.concl_of arg_cong;
+ val rep_const = Thm.cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
+ val cong' = cterm_instantiate [(Thm.cterm_of thy cong_f, rep_const)] arg_cong;
+ val dist = cterm_instantiate [(Thm.cterm_of thy distinct_f, rep_const)] distinct_lemma;
val (thy', defs', eqns', _) =
fold (make_constr_def typedef T (length constrs))
(constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
@@ -390,7 +390,7 @@
(* prove inj Rep_dt_i and Rep_dt_i x : rep_set_dt_i *)
val fun_congs =
- map (fn T => make_elim (Drule.instantiate' [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
+ map (fn T => make_elim (Drule.instantiate' [SOME (Thm.ctyp_of thy5 T)] [] fun_cong)) branchTs;
fun prove_iso_thms ds (inj_thms, elem_thms) =
let
@@ -609,7 +609,7 @@
val (indrule_lemma_prems, indrule_lemma_concls) =
split_list (map2 mk_indrule_lemma descr' recTs);
- val cert = cterm_of thy6;
+ val cert = Thm.cterm_of thy6;
val indrule_lemma =
Goal.prove_sorry_global thy6 [] []
@@ -623,7 +623,7 @@
[TRY (rtac conjI 1), resolve_tac ctxt Rep_inverse_thms 1,
etac mp 1, resolve_tac ctxt iso_elem_thms 1])]);
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma)));
val frees =
if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))]
else map (Free o apfst fst o dest_Var) Ps;
--- a/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Wed Mar 04 19:53:18 2015 +0100
@@ -126,9 +126,9 @@
fun ind_tac indrule indnames = CSUBGOAL (fn (cgoal, i) =>
let
- val cert = cterm_of (Thm.theory_of_cterm cgoal);
- val goal = term_of cgoal;
- val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
+ val cert = Thm.cterm_of (Thm.theory_of_cterm cgoal);
+ val goal = Thm.term_of cgoal;
+ val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule));
val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop (Logic.strip_imp_concl goal));
val getP =
if can HOLogic.dest_imp (hd ts)
@@ -158,16 +158,18 @@
fun exh_tac ctxt exh_thm_of = CSUBGOAL (fn (cgoal, i) =>
let
val thy = Thm.theory_of_cterm cgoal;
- val goal = term_of cgoal;
+ val goal = Thm.term_of cgoal;
val params = Logic.strip_params goal;
val (_, Type (tname, _)) = hd (rev params);
val exhaustion = Thm.lift_rule cgoal (exh_thm_of tname);
- val prem' = hd (prems_of exhaustion);
+ val prem' = hd (Thm.prems_of exhaustion);
val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
val exhaustion' =
- cterm_instantiate [(cterm_of thy (head_of lhs),
- cterm_of thy (fold_rev (fn (_, T) => fn t => Abs ("z", T, t)) params (Bound 0)))] exhaustion;
- in compose_tac ctxt (false, exhaustion', nprems_of exhaustion) i end);
+ cterm_instantiate
+ [(Thm.cterm_of thy (head_of lhs),
+ Thm.cterm_of thy (fold_rev (fn (_, T) => fn t => Abs ("z", T, t)) params (Bound 0)))]
+ exhaustion;
+ in compose_tac ctxt (false, exhaustion', Thm.nprems_of exhaustion) i end);
(********************** Internal description of datatypes *********************)
--- a/src/HOL/Tools/Old_Datatype/old_datatype_data.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype_data.ML Wed Mar 04 19:53:18 2015 +0100
@@ -87,12 +87,12 @@
val info_of_case = Symtab.lookup o #cases o Data.get;
fun ctrs_of_exhaust exhaust =
- Logic.strip_imp_prems (prop_of exhaust) |>
+ Logic.strip_imp_prems (Thm.prop_of exhaust) |>
map (head_of o snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o the_single
o Logic.strip_assums_hyp);
fun case_of_case_rewrite case_rewrite =
- head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of case_rewrite))));
+ head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of case_rewrite))));
fun ctr_sugar_of_info ({exhaust, nchotomy, inject, distinct, case_rewrites, case_cong,
case_cong_weak, split, split_asm, ...} : Old_Datatype_Aux.info) =
--- a/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Wed Mar 04 19:53:18 2015 +0100
@@ -35,7 +35,8 @@
val newTs = take (length (hd descr)) recTs;
val maxidx = Thm.maxidx_of induct;
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+ val induct_Ps =
+ map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)));
fun prove_casedist_thm (i, (T, t)) =
let
@@ -45,7 +46,7 @@
Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx + 1), T), Bound 0) $
Var (("P", 0), HOLogic.boolT));
val insts = take i dummyPs @ (P :: drop (i + 1) dummyPs);
- val cert = cterm_of thy;
+ val cert = Thm.cterm_of thy;
val insts' = map cert induct_Ps ~~ map cert insts;
val induct' =
refl RS
@@ -87,7 +88,8 @@
val used = fold Term.add_tfree_namesT recTs [];
val newTs = take (length (hd descr)) recTs;
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+ val induct_Ps =
+ map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)));
val big_rec_name' = "rec_set_" ^ big_name;
val rec_set_names' =
@@ -202,7 +204,7 @@
Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
absfree ("y", T2) (set_t $ Old_Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
(rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
- val cert = cterm_of thy1;
+ val cert = Thm.cterm_of thy1;
val insts =
map (fn ((i, T), t) => absfree ("x" ^ string_of_int i, T) t)
((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
@@ -351,7 +353,7 @@
map2 prove_cases newTs (Old_Datatype_Prop.make_cases case_names0 descr thy2);
fun case_name_of (th :: _) =
- fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))));
+ fst (dest_Const (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th))))));
val case_names = map case_name_of case_thms;
in
@@ -377,8 +379,8 @@
fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) =
let
- val cert = cterm_of thy;
- val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
+ val cert = Thm.cterm_of thy;
+ val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (Thm.prems_of exhaustion)));
val exhaustion' = cterm_instantiate [(cert lhs, cert (Free ("x", T)))] exhaustion;
fun tac ctxt =
EVERY [resolve_tac ctxt [exhaustion'] 1,
@@ -448,9 +450,9 @@
let
val Const (@{const_name Pure.imp}, _) $ tm $ _ = t;
val Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma) = tm;
- val cert = cterm_of thy;
+ val cert = Thm.cterm_of thy;
val nchotomy' = nchotomy RS spec;
- val [v] = Term.add_vars (concl_of nchotomy') [];
+ val [v] = Term.add_vars (Thm.concl_of nchotomy') [];
val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy';
in
Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML Wed Mar 04 19:53:18 2015 +0100
@@ -319,7 +319,7 @@
fun translate_intros ensure_groundness ctxt gr const constant_table =
let
val intros = map (preprocess_intro (Proof_Context.theory_of ctxt)) (Graph.get_node gr const)
- val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
+ val (intros', ctxt') = Variable.import_terms true (map Thm.prop_of intros) ctxt
val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
fun translate_intro intro =
let
--- a/src/HOL/Tools/Predicate_Compile/core_data.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/core_data.ML Wed Mar 04 19:53:18 2015 +0100
@@ -235,7 +235,7 @@
fun term_pair_of (ix, (ty, t)) = (Var (ix, ty), t)
fun inst_of_matches tts =
fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty)
- |> snd |> Vartab.dest |> map (apply2 (cterm_of thy) o term_pair_of)
+ |> snd |> Vartab.dest |> map (apply2 (Thm.cterm_of thy) o term_pair_of)
val (cases, (eqs, prems1)) = apsnd (chop (nargs - nparams)) (chop n prems)
val case_th =
rewrite_rule ctxt2 (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1))
@@ -247,8 +247,8 @@
Thm.instantiate ([], inst_of_matches pats) case_th
OF replicate nargs @{thm refl}
val thesis =
- Thm.instantiate ([], inst_of_matches (prems_of case_th' ~~ map prop_of prems2)) case_th'
- OF prems2
+ Thm.instantiate ([], inst_of_matches (Thm.prems_of case_th' ~~ map Thm.prop_of prems2))
+ case_th' OF prems2
in rtac thesis 1 end
in
Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule
@@ -277,7 +277,7 @@
val pos = Position.thread_data ()
fun is_intro_of intro =
let
- val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
+ val (const, _) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of intro))
in (fst (dest_Const const) = name) end;
val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result))
val index = find_index (fn s => s = name) (#names (fst info))
@@ -330,7 +330,7 @@
fun set_elim thm =
let
val (name, _) =
- dest_Const (fst (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
+ dest_Const (fst (strip_comb (HOLogic.dest_Trueprop (hd (Thm.prems_of thm)))))
in
PredData.map (Graph.map_node name (map_pred_data (apsnd (apfst (apfst (apsnd (K (SOME thm))))))))
end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Wed Mar 04 19:53:18 2015 +0100
@@ -459,13 +459,13 @@
(Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) = true
| is_equationlike_term _ = false
-val is_equationlike = is_equationlike_term o prop_of
+val is_equationlike = is_equationlike_term o Thm.prop_of
fun is_pred_equation_term (Const (@{const_name Pure.eq}, _) $ u $ v) =
(fastype_of u = @{typ bool}) andalso (fastype_of v = @{typ bool})
| is_pred_equation_term _ = false
-val is_pred_equation = is_pred_equation_term o prop_of
+val is_pred_equation = is_pred_equation_term o Thm.prop_of
fun is_intro_term constname t =
the_default false (try (fn t =>
@@ -473,7 +473,7 @@
Const (c, _) => c = constname
| _ => false) t)
-fun is_intro constname t = is_intro_term constname (prop_of t)
+fun is_intro constname t = is_intro_term constname (Thm.prop_of t)
fun is_predT (T as Type("fun", [_, _])) = (body_type T = @{typ bool})
| is_predT _ = false
@@ -528,7 +528,8 @@
val t'' = Term.subst_bounds (rev vs, t');
in ((ps', t''), nctxt') end
-val strip_intro_concl = strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of
+val strip_intro_concl =
+ strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of
(* introduction rule combinators *)
@@ -610,7 +611,7 @@
(* lifting term operations to theorems *)
fun map_term thy f th =
- Skip_Proof.make_thm thy (f (prop_of th))
+ Skip_Proof.make_thm thy (f (Thm.prop_of th))
(*
fun equals_conv lhs_cv rhs_cv ct =
@@ -860,23 +861,23 @@
end
fun dest_conjunct_prem th =
- (case HOLogic.dest_Trueprop (prop_of th) of
+ (case HOLogic.dest_Trueprop (Thm.prop_of th) of
(Const (@{const_name HOL.conj}, _) $ _ $ _) =>
- dest_conjunct_prem (th RS @{thm conjunct1})
- @ dest_conjunct_prem (th RS @{thm conjunct2})
- | _ => [th])
+ dest_conjunct_prem (th RS @{thm conjunct1}) @
+ dest_conjunct_prem (th RS @{thm conjunct2})
+ | _ => [th])
fun expand_tuples thy intro =
let
val ctxt = Proof_Context.init_global thy
val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
- val intro_t = prop_of intro'
+ val intro_t = Thm.prop_of intro'
val concl = Logic.strip_imp_concl intro_t
val (_, args) = strip_comb (HOLogic.dest_Trueprop concl)
val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
val (pats', _, ctxt3) = fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
fun rewrite_pat (ct1, ct2) =
- (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
+ (ct1, Thm.cterm_of thy (Pattern.rewrite_term thy pats' [] (Thm.term_of ct2)))
val t_insts' = map rewrite_pat t_insts
val intro'' = Thm.instantiate (T_insts, t_insts') intro
val [intro'''] = Variable.export ctxt3 ctxt [intro'']
@@ -939,7 +940,7 @@
let
val th = case_rewrite thy Tcon
val ctxt = Proof_Context.init_global thy
- val f = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th)))))
+ val f = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th)))))
val Type ("fun", [uninst_T, uninst_T']) = fastype_of f
val ([yname], ctxt') = Variable.add_fixes ["y"] ctxt
val T' = TFree ("'t'", @{sort type})
@@ -1026,7 +1027,7 @@
(* Some last processing *)
fun remove_pointless_clauses intro =
- if Logic.strip_imp_prems (prop_of intro) = [@{prop "False"}] then
+ if Logic.strip_imp_prems (Thm.prop_of intro) = [@{prop "False"}] then
[]
else [intro]
@@ -1045,7 +1046,7 @@
map_filter_premises (fn p => if p = @{prop True} then NONE else SOME p) intro_t
in
Option.map (Skip_Proof.make_thm thy)
- (process_False (process_True (prop_of (process intro))))
+ (process_False (process_True (Thm.prop_of (process intro))))
end
@@ -1089,7 +1090,7 @@
fun instantiate_ho_args th =
let
val (_, args') =
- (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
+ (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of) th
val ho_args' = map dest_Var (ho_args_of_typ T args')
in Thm.certify_instantiate ([], ho_args' ~~ ho_args) th end
val outp_pred =
@@ -1136,7 +1137,7 @@
val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
(* TODO: distinct required ? -- test case with more than one parameter! *)
val params = distinct (op aconv) params
- val intros = map prop_of intros_th
+ val intros = map Thm.prop_of intros_th
val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
val argsT = binder_types (fastype_of pred)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1226,7 +1226,7 @@
fun prepare_intrs options ctxt prednames intros =
let
val thy = Proof_Context.theory_of ctxt
- val intrs = map prop_of intros
+ val intrs = map Thm.prop_of intros
val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
val (preds, intrs) = unify_consts thy preds intrs
val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] ctxt
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Mar 04 19:53:18 2015 +0100
@@ -72,17 +72,17 @@
| _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t])
end
*)
-val defining_term_of_introrule = defining_term_of_introrule_term o prop_of
+val defining_term_of_introrule = defining_term_of_introrule_term o Thm.prop_of
fun defining_const_of_introrule th =
(case defining_term_of_introrule th of
Const (c, _) => c
- | _ => raise TERM ("defining_const_of_introrule failed: Not a constant", [prop_of th]))
+ | _ => raise TERM ("defining_const_of_introrule failed: Not a constant", [Thm.prop_of th]))
(*TODO*)
fun is_introlike_term _ = true
-val is_introlike = is_introlike_term o prop_of
+val is_introlike = is_introlike_term o Thm.prop_of
fun check_equation_format_term (t as (Const (@{const_name Pure.eq}, _) $ u $ _)) =
(case strip_comb u of
@@ -95,19 +95,19 @@
| check_equation_format_term t =
raise TERM ("check_equation_format_term failed: Not an equation", [t])
-val check_equation_format = check_equation_format_term o prop_of
+val check_equation_format = check_equation_format_term o Thm.prop_of
fun defining_term_of_equation_term (Const (@{const_name Pure.eq}, _) $ u $ _) = fst (strip_comb u)
| defining_term_of_equation_term t =
raise TERM ("defining_const_of_equation_term failed: Not an equation", [t])
-val defining_term_of_equation = defining_term_of_equation_term o prop_of
+val defining_term_of_equation = defining_term_of_equation_term o Thm.prop_of
fun defining_const_of_equation th =
(case defining_term_of_equation th of
Const (c, _) => c
- | _ => raise TERM ("defining_const_of_equation failed: Not a constant", [prop_of th]))
+ | _ => raise TERM ("defining_const_of_equation failed: Not a constant", [Thm.prop_of th]))
@@ -115,7 +115,7 @@
(* Normalizing equations *)
fun mk_meta_equation th =
- (case prop_of th of
+ (case Thm.prop_of th of
Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _) =>
th RS @{thm eq_reflection}
| _ => th)
@@ -124,7 +124,7 @@
fun full_fun_cong_expand th =
let
- val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
+ val (f, args) = strip_comb (fst (Logic.dest_equals (Thm.prop_of th)))
val i = length (binder_types (fastype_of f)) - length args
in funpow i (fn th => th RS meta_fun_cong) th end;
@@ -137,7 +137,7 @@
let
val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *)
val ((_, [th']), _) = Variable.import true [th] ctxt
- val t = prop_of th'
+ val t = Thm.prop_of th'
val frees = Term.add_frees t []
val freenames = Term.add_free_names t []
val nctxt = Name.make_context freenames
@@ -269,7 +269,7 @@
Type (Tcon, _) => can (Ctr_Sugar.dest_ctr ctxt Tcon) (Const x)
| _ => false)
fun defiants_of specs =
- fold (Term.add_consts o prop_of) specs []
+ fold (Term.add_consts o Thm.prop_of) specs []
|> filter_out is_datatype_constructor
|> filter_out is_nondefining_const
|> filter_out has_code_pred_intros
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Wed Mar 04 19:53:18 2015 +0100
@@ -169,7 +169,7 @@
SOME raw_split_thm =>
let
val split_thm = prepare_split_thm (Proof_Context.init_global thy) raw_split_thm
- val (assms, concl) = Logic.strip_horn (prop_of split_thm)
+ val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
val (_, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl)
val t' = case_betapply thy t
val subst = Pattern.match thy (split_t, t') (Vartab.empty, Vartab.empty)
@@ -311,7 +311,7 @@
let
fun lookup_pred t = lookup thy (Fun_Pred.get thy) t
(*val _ = tracing ("Rewriting intro " ^ Display.string_of_thm_global thy intro)*)
- val intro_t = Logic.unvarify_global (prop_of intro)
+ val intro_t = Logic.unvarify_global (Thm.prop_of intro)
val (prems, concl) = Logic.strip_horn intro_t
val frees = map fst (Term.add_frees intro_t [])
fun rewrite prem names =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Wed Mar 04 19:53:18 2015 +0100
@@ -36,7 +36,7 @@
(* TODO: contextify things - this line is to unvarify the split_thm *)
(*val ((_, [isplit_thm]), _) =
Variable.import true [split_thm] (Proof_Context.init_global thy)*)
- val (assms, concl) = Logic.strip_horn (prop_of split_thm)
+ val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
val (_, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl)
val atom' = case_betapply thy atom
val subst = Pattern.match thy (split_t, atom') (Vartab.empty, Vartab.empty)
@@ -132,7 +132,7 @@
val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *)
val ((_, intros), ctxt') = Variable.import true intros ctxt
val (intros', (local_defs, thy')) = (fold_map o fold_map_atoms)
- (flatten constname) (map prop_of intros) ([], thy)
+ (flatten constname) (map Thm.prop_of intros) ([], thy)
val ctxt'' = Proof_Context.transfer thy' ctxt'
val intros'' =
map (fn t => Goal.prove ctxt'' [] [] t (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt''))) intros'
@@ -147,7 +147,7 @@
val ((_, ths'), ctxt') = Variable.import true ths ctxt
fun introrulify' th =
let
- val (lhs, rhs) = Logic.dest_equals (prop_of th)
+ val (lhs, rhs) = Logic.dest_equals (Thm.prop_of th)
val frees = Term.add_free_names rhs []
val disjuncts = HOLogic.dest_disj rhs
val nctxt = Name.make_context frees
@@ -158,14 +158,14 @@
in
(ps, Logic.list_implies (prems, HOLogic.mk_Trueprop lhs))
end
- val x = ((cterm_of thy) o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o
- Logic.dest_implies o prop_of) @{thm exI}
+ val x = (Thm.cterm_of thy o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o
+ Logic.dest_implies o Thm.prop_of) @{thm exI}
fun prove_introrule (index, (ps, introrule)) =
let
val tac = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [th]) 1
THEN Inductive.select_disj_tac ctxt' (length disjuncts) (index + 1) 1
THEN (EVERY (map (fn y =>
- rtac (Drule.cterm_instantiate [(x, cterm_of thy (Free y))] @{thm exI}) 1) ps))
+ rtac (Drule.cterm_instantiate [(x, Thm.cterm_of thy (Free y))] @{thm exI}) 1) ps))
THEN REPEAT_DETERM (rtac @{thm conjI} 1 THEN assume_tac ctxt' 1)
THEN TRY (assume_tac ctxt' 1)
in
@@ -270,9 +270,9 @@
fun flat_intro intro (new_defs, thy) =
let
val constname = fst (dest_Const (fst (strip_comb
- (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of intro))))))
+ (HOLogic.dest_Trueprop (Logic.strip_imp_concl (Thm.prop_of intro))))))
val (intro_ts, (new_defs, thy)) =
- fold_map_atoms (process constname) (prop_of intro) (new_defs, thy)
+ fold_map_atoms (process constname) (Thm.prop_of intro) (new_defs, thy)
val th = Skip_Proof.make_thm thy intro_ts
in
(th, (new_defs, thy))
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Wed Mar 04 19:53:18 2015 +0100
@@ -108,10 +108,10 @@
THEN Subgoal.FOCUS_PREMS (fn {context = ctxt', prems, concl, ...} =>
let
val param_prem = nth prems premposition
- val (param, _) = strip_comb (HOLogic.dest_Trueprop (prop_of param_prem))
+ val (param, _) = strip_comb (HOLogic.dest_Trueprop (Thm.prop_of param_prem))
val prems' = maps dest_conjunct_prem (take nargs prems)
fun param_rewrite prem =
- param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of prem)))
+ param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of prem)))
val SOME rew_eq = find_first param_rewrite prems'
val param_prem' = rewrite_rule ctxt'
(map (fn th => th RS @{thm eq_reflection})
@@ -122,7 +122,7 @@
end) ctxt 1
THEN trace_tac ctxt options "after prove parameter call")
-fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st
+fun SOLVED tac st = FILTER (fn st' => Thm.nprems_of st' = Thm.nprems_of st - 1) tac st
fun prove_match options ctxt nargs out_ts =
let
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Wed Mar 04 19:53:18 2015 +0100
@@ -31,7 +31,7 @@
let
val ((_, intros'), ctxt') = Variable.importT intros ctxt
val pred' =
- fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of (hd intros')))))
+ fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl (Thm.prop_of (hd intros')))))
val Ts = binder_types (fastype_of pred')
val argTs = map fastype_of args
val Tsubst = Type.raw_matches (argTs, Ts) Vartab.empty
@@ -61,8 +61,8 @@
fun specialise_intros black_list (pred, intros) pats thy =
let
val ctxt = Proof_Context.init_global thy
- val maxidx = fold (Term.maxidx_term o prop_of) intros ~1
- val pats = map (Logic.incr_indexes ([], maxidx + 1)) pats
+ val maxidx = fold (Term.maxidx_term o Thm.prop_of) intros ~1
+ val pats = map (Logic.incr_indexes ([], maxidx + 1)) pats
val (((pred, intros), pats), ctxt') = import (pred, intros) pats ctxt
val result_pats = map Var (fold_rev Term.add_vars pats [])
fun mk_fresh_name names =
@@ -82,7 +82,7 @@
val specialised_const = Const (constname, constT)
fun specialise_intro intro =
(let
- val (prems, concl) = Logic.strip_horn (prop_of intro)
+ val (prems, concl) = Logic.strip_horn (Thm.prop_of intro)
val env = Pattern.unify (Context.Theory thy)
(HOLogic.mk_Trueprop (list_comb (pred, pats)), concl) (Envir.empty 0)
val prems = map (Envir.norm_term env) prems
@@ -201,7 +201,7 @@
let
(* FIXME: only necessary because of sloppy Logic.unvarify in restrict_pattern *)
val intros = Drule.zero_var_indexes_list intros
- val (intros_t', thy') = (fold_map o fold_map_atoms) detect' (map prop_of intros) thy
+ val (intros_t', thy') = (fold_map o fold_map_atoms) detect' (map Thm.prop_of intros) thy
in
((constname, map (Skip_Proof.make_thm thy') intros_t'), thy')
end
--- a/src/HOL/Tools/Qelim/cooper.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Wed Mar 04 19:53:18 2015 +0100
@@ -122,22 +122,22 @@
| Dvd of cterm*cterm | NDvd of cterm*cterm | Nox
fun whatis x ct =
-( case (term_of ct) of
+( case Thm.term_of ct of
Const(@{const_name HOL.conj},_)$_$_ => And (Thm.dest_binop ct)
| Const (@{const_name HOL.disj},_)$_$_ => Or (Thm.dest_binop ct)
-| Const (@{const_name HOL.eq},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
+| Const (@{const_name HOL.eq},_)$y$_ => if Thm.term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
| Const (@{const_name Not},_) $ (Const (@{const_name HOL.eq},_)$y$_) =>
- if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
+ if Thm.term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
| Const (@{const_name Orderings.less}, _) $ y$ z =>
- if term_of x aconv y then Lt (Thm.dest_arg ct)
- else if term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox
+ if Thm.term_of x aconv y then Lt (Thm.dest_arg ct)
+ else if Thm.term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox
| Const (@{const_name Orderings.less_eq}, _) $ y $ z =>
- if term_of x aconv y then Le (Thm.dest_arg ct)
- else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox
+ if Thm.term_of x aconv y then Le (Thm.dest_arg ct)
+ else if Thm.term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox
| Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_) =>
- if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox
+ if Thm.term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox
| Const (@{const_name Not},_) $ (Const (@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$y$_)) =>
- if term_of x aconv y then
+ if Thm.term_of x aconv y then
NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox
| _ => Nox)
handle CTERM _ => Nox;
@@ -148,7 +148,7 @@
(Thm.dest_arg t)
in (Thm.lambda x o Thm.dest_arg o Thm.dest_arg) eq end;
-val get_pmi = get_pmi_term o cprop_of;
+val get_pmi = get_pmi_term o Thm.cprop_of;
val p_v' = @{cpat "?P' :: int => bool"};
val q_v' = @{cpat "?Q' :: int => bool"};
@@ -174,7 +174,7 @@
val cmulC = @{cterm "op * :: int => _"}
val cminus = @{cterm "op - :: int => _"}
val cone = @{cterm "1 :: int"}
-val [addC, mulC, subC] = map term_of [cadd, cmulC, cminus]
+val [addC, mulC, subC] = map Thm.term_of [cadd, cmulC, cminus]
val [zero, one] = [@{term "0 :: int"}, @{term "1 :: int"}];
fun numeral1 f n = HOLogic.mk_number iT (f (dest_number n));
@@ -294,7 +294,7 @@
| lin vs fm = fm;
fun lint_conv ctxt vs ct =
-let val t = term_of ct
+let val t = Thm.term_of ct
in (provelin ctxt ((HOLogic.eq_const iT)$t$(lint vs t) |> HOLogic.mk_Trueprop))
RS eq_reflection
end;
@@ -305,18 +305,18 @@
| is_intrel (@{term "Not"}$(b$_$_)) = is_intrel_type (fastype_of b)
| is_intrel _ = false;
-fun linearize_conv ctxt vs ct = case term_of ct of
+fun linearize_conv ctxt vs ct = case Thm.term_of ct of
Const(@{const_name Rings.dvd},_)$_$_ =>
let
val th = Conv.binop_conv (lint_conv ctxt vs) ct
val (d',t') = Thm.dest_binop (Thm.rhs_of th)
- val (dt',tt') = (term_of d', term_of t')
+ val (dt',tt') = (Thm.term_of d', Thm.term_of t')
in if is_number dt' andalso is_number tt'
then Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset presburger_ss ctxt))) th
else
let
val dth =
- case perhaps_number (term_of d') of
+ case perhaps_number (Thm.term_of d') of
SOME d => if d < 0 then
(Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (lint_conv ctxt vs)))
(Thm.transitive th (inst' [d',t'] dvd_uminus))
@@ -345,10 +345,10 @@
fun unify ctxt q =
let
val (e,(cx,p)) = q |> Thm.dest_comb ||> Thm.dest_abs NONE
- val x = term_of cx
+ val x = Thm.term_of cx
val ins = insert (op = : int * int -> bool)
fun h (acc,dacc) t =
- case (term_of t) of
+ case Thm.term_of t of
Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
if x aconv y andalso member (op =)
[@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
@@ -366,7 +366,7 @@
val (cs,ds) = h ([],[]) p
val l = Integer.lcms (union (op =) cs ds)
fun cv k ct =
- let val (tm as b$s$t) = term_of ct
+ let val (tm as b$s$t) = Thm.term_of ct
in ((HOLogic.eq_const bT)$tm$(b$(linear_cmul k s)$(linear_cmul k t))
|> HOLogic.mk_Trueprop |> provelin ctxt) RS eq_reflection end
fun nzprop x =
@@ -381,13 +381,13 @@
let val tab = fold Inttab.update
(ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty
in
- fn ct => the (Inttab.lookup tab (ct |> term_of |> dest_number))
+ fn ct => the (Inttab.lookup tab (ct |> Thm.term_of |> dest_number))
handle Option.Option =>
(writeln ("noz: Theorems-Table contains no entry for " ^
Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option.Option)
end
fun unit_conv t =
- case (term_of t) of
+ case Thm.term_of t of
Const(@{const_name HOL.conj},_)$_$_ => Conv.binop_conv unit_conv t
| Const(@{const_name HOL.disj},_)$_$_ => Conv.binop_conv unit_conv t
| Const (@{const_name Not},_)$_ => Conv.arg_conv unit_conv t
@@ -420,8 +420,8 @@
val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth)
val th' = inst' [Thm.lambda ltx (Thm.rhs_of uth), clt] unity_coeff_ex
val thf = Thm.transitive th
- (Thm.transitive (Thm.symmetric (Thm.beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th')
- val (lth,rth) = Thm.dest_comb (cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true
+ (Thm.transitive (Thm.symmetric (Thm.beta_conversion true (Thm.cprop_of th' |> Thm.dest_arg1))) th')
+ val (lth,rth) = Thm.dest_comb (Thm.cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true
||> Thm.beta_conversion true |>> Thm.symmetric
in Thm.transitive (Thm.transitive lth thf) rth end;
@@ -430,7 +430,7 @@
val insert_tm = @{cterm "insert :: int => _"};
fun mkISet cts = fold_rev (Thm.apply insert_tm #> Thm.apply) cts emptyIS;
val eqelem_imp_imp = @{thm eqelem_imp_iff} RS iffD1;
-val [A_tm,B_tm] = map (fn th => cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg
+val [A_tm,B_tm] = map (fn th => Thm.cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg
|> Thm.dest_abs NONE |> snd |> Thm.dest_fun |> Thm.dest_arg)
[asetP,bsetP];
@@ -454,8 +454,8 @@
| Le t => (bacc, ins (plus1 t) aacc,dacc)
| Gt t => (ins t bacc, aacc,dacc)
| Ge t => (ins (minus1 t) bacc, aacc,dacc)
- | Dvd (d,_) => (bacc,aacc,insert (op =) (term_of d |> dest_number) dacc)
- | NDvd (d,_) => (bacc,aacc,insert (op =) (term_of d|> dest_number) dacc)
+ | Dvd (d,_) => (bacc,aacc,insert (op =) (Thm.term_of d |> dest_number) dacc)
+ | NDvd (d,_) => (bacc,aacc,insert (op =) (Thm.term_of d|> dest_number) dacc)
| _ => (bacc, aacc, dacc)
val (b0,a0,ds) = h p ([],[],[])
val d = Integer.lcms ds
@@ -469,7 +469,7 @@
in Thm.equal_elim (Thm.symmetric th) TrueI end;
val dvd =
let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in
- fn ct => the (Inttab.lookup tab (term_of ct |> dest_number))
+ fn ct => the (Inttab.lookup tab (Thm.term_of ct |> dest_number))
handle Option.Option =>
(writeln ("dvd: Theorems-Table contains no entry for" ^
Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option.Option)
@@ -485,18 +485,18 @@
val insI2 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI2"}
in
fun provein x S =
- case term_of S of
+ case Thm.term_of S of
Const(@{const_name Orderings.bot}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
| Const(@{const_name insert}, _) $ y $ _ =>
let val (cy,S') = Thm.dest_binop S
- in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
+ in if Thm.term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
else Thm.implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
(provein x S')
end
end
- val al = map (lint vs o term_of) a0
- val bl = map (lint vs o term_of) b0
+ val al = map (lint vs o Thm.term_of) a0
+ val bl = map (lint vs o Thm.term_of) b0
val (sl,s0,f,abths,cpth) =
if length (distinct (op aconv) bl) <= length (distinct (op aconv) al)
then
@@ -516,29 +516,29 @@
val cpth =
let
val sths = map (fn (tl,t0) =>
- if tl = term_of t0
+ if tl = Thm.term_of t0
then instantiate' [SOME @{ctyp "int"}] [SOME t0] refl
- else provelin ctxt ((HOLogic.eq_const iT)$tl$(term_of t0)
+ else provelin ctxt ((HOLogic.eq_const iT)$tl$(Thm.term_of t0)
|> HOLogic.mk_Trueprop))
(sl ~~ s0)
- val csl = distinct (op aconvc) (map (cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths)
+ val csl = distinct (op aconvc) (map (Thm.cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths)
val S = mkISet csl
- val inStab = fold (fn ct => fn tab => Termtab.update (term_of ct, provein ct S) tab)
+ val inStab = fold (fn ct => fn tab => Termtab.update (Thm.term_of ct, provein ct S) tab)
csl Termtab.empty
val eqelem_th = instantiate' [SOME @{ctyp "int"}] [NONE,NONE, SOME S] eqelem_imp_imp
val inS =
let
val tab = fold Termtab.update
(map (fn eq =>
- let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop
- val th = if term_of s = term_of t
- then the (Termtab.lookup inStab (term_of s))
+ let val (s,t) = Thm.cprop_of eq |> Thm.dest_arg |> Thm.dest_binop
+ val th = if Thm.term_of s = Thm.term_of t (* FIXME equality? *)
+ then the (Termtab.lookup inStab (Thm.term_of s))
else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)
- [eq, the (Termtab.lookup inStab (term_of s))]
- in (term_of t, th) end)
+ [eq, the (Termtab.lookup inStab (Thm.term_of s))]
+ in (Thm.term_of t, th) end)
sths) Termtab.empty
in
- fn ct => the (Termtab.lookup tab (term_of ct))
+ fn ct => the (Termtab.lookup tab (Thm.term_of ct))
handle Option.Option =>
(writeln ("inS: No theorem for " ^ Syntax.string_of_term ctxt (Thm.term_of ct));
raise Option.Option)
@@ -552,7 +552,7 @@
fun literals_conv bops uops env cv =
let fun h t =
- case (term_of t) of
+ case Thm.term_of t of
b$_$_ => if member (op aconv) bops b then Conv.binop_conv h t else cv env t
| u$_ => if member (op aconv) uops u then Conv.arg_conv h t else cv env t
| _ => cv env t
@@ -571,7 +571,7 @@
(Simplifier.rewrite (put_simpset conv_ss ctxt))
(Simplifier.rewrite (put_simpset presburger_ss ctxt))
(Simplifier.rewrite (put_simpset conv_ss ctxt))
- (cons o term_of) (Misc_Legacy.term_frees (term_of p))
+ (cons o Thm.term_of) (Misc_Legacy.term_frees (Thm.term_of p))
(linearize_conv ctxt) (integer_nnf_conv ctxt)
(cooperex_conv ctxt) p
handle CTERM _ => raise COOPER "bad cterm"
@@ -708,7 +708,7 @@
| _ => [ct]);
fun strip_objall ct =
- case term_of ct of
+ case Thm.term_of ct of
Const (@{const_name All}, _) $ Abs (xn,_,_) =>
let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct
in apfst (cons (a,v)) (strip_objall t')
@@ -760,13 +760,14 @@
| _ => is_number t orelse can HOLogic.dest_nat t
fun ty cts t =
- if not (member (op =) [HOLogic.intT, HOLogic.natT, HOLogic.boolT] (typ_of (ctyp_of_term t))) then false
- else case term_of t of
- c$l$r => if member (op =) [@{term"op *::int => _"}, @{term"op *::nat => _"}] c
- then not (isnum l orelse isnum r)
- else not (member (op aconv) cts c)
- | c$_ => not (member (op aconv) cts c)
- | c => not (member (op aconv) cts c)
+ if not (member (op =) [HOLogic.intT, HOLogic.natT, HOLogic.boolT] (Thm.typ_of (Thm.ctyp_of_term t)))
+ then false
+ else case Thm.term_of t of
+ c$l$r => if member (op =) [@{term"op *::int => _"}, @{term"op *::nat => _"}] c
+ then not (isnum l orelse isnum r)
+ else not (member (op aconv) cts c)
+ | c$_ => not (member (op aconv) cts c)
+ | c => not (member (op aconv) cts c)
val term_constants =
let fun h acc t = case t of
@@ -777,15 +778,19 @@
in h [] end;
in
fun is_relevant ctxt ct =
- subset (op aconv) (term_constants (term_of ct) , snd (get ctxt))
- andalso forall (fn Free (_,T) => member (op =) [@{typ int}, @{typ nat}] T) (Misc_Legacy.term_frees (term_of ct))
- andalso forall (fn Var (_,T) => member (op =) [@{typ int}, @{typ nat}] T) (Misc_Legacy.term_vars (term_of ct));
+ subset (op aconv) (term_constants (Thm.term_of ct), snd (get ctxt))
+ andalso
+ forall (fn Free (_, T) => member (op =) [@{typ int}, @{typ nat}] T)
+ (Misc_Legacy.term_frees (Thm.term_of ct))
+ andalso
+ forall (fn Var (_, T) => member (op =) [@{typ int}, @{typ nat}] T)
+ (Misc_Legacy.term_vars (Thm.term_of ct));
fun int_nat_terms ctxt ct =
let
val cts = snd (get ctxt)
fun h acc t = if ty cts t then insert (op aconvc) t acc else
- case (term_of t) of
+ case Thm.term_of t of
_$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t)
| Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc))
| _ => acc
@@ -795,8 +800,8 @@
fun generalize_tac f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st =>
let
fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "Pure.all"}
- fun gen x t = Thm.apply (all (ctyp_of_term x)) (Thm.lambda x t)
- val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p)
+ fun gen x t = Thm.apply (all (Thm.ctyp_of_term x)) (Thm.lambda x t)
+ val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (Thm.term_of a, Thm.term_of b)) (f p)
val p' = fold_rev gen ts p
in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end));
@@ -849,7 +854,7 @@
let
val cpth =
if Config.get ctxt quick_and_dirty
- then oracle (ctxt, Envir.beta_norm (Envir.eta_long [] (term_of (Thm.dest_arg p))))
+ then oracle (ctxt, Envir.beta_norm (Envir.eta_long [] (Thm.term_of (Thm.dest_arg p))))
else Conv.arg_conv (conv ctxt) p
val p' = Thm.rhs_of cpth
val th = Thm.implies_intr p' (Thm.equal_elim (Thm.symmetric cpth) (Thm.assume p'))
@@ -891,7 +896,7 @@
val constsN = "consts";
val any_keyword = keyword constsN
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-val terms = thms >> map (term_of o Drule.dest_term);
+val terms = thms >> map (Thm.term_of o Drule.dest_term);
fun optional scan = Scan.optional scan [];
--- a/src/HOL/Tools/Qelim/qelim.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Qelim/qelim.ML Wed Mar 04 19:53:18 2015 +0100
@@ -21,7 +21,7 @@
fun gen_qelim_conv precv postcv simpex_conv ins env atcv ncv qcv =
let
fun conv env p =
- case (term_of p) of
+ case Thm.term_of p of
Const(s,T)$_$_ =>
if domain_type T = HOLogic.boolT
andalso member (op =) [@{const_name HOL.conj}, @{const_name HOL.disj},
@@ -37,7 +37,7 @@
val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p')
|> Drule.arg_cong_rule e
val th' = simpex_conv (Thm.rhs_of th)
- val (_, r) = Thm.dest_equals (cprop_of th')
+ val (_, r) = Thm.dest_equals (Thm.cprop_of th')
in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
else Thm.transitive (Thm.transitive th th') (conv env r) end
| Const(@{const_name Ex},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
--- a/src/HOL/Tools/Quickcheck/find_unused_assms.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quickcheck/find_unused_assms.ML Wed Mar 04 19:53:18 2015 +0100
@@ -14,7 +14,7 @@
struct
fun thms_of thy thy_name = Global_Theory.all_thms_of thy false
- |> filter (fn (_, th) => Context.theory_name (theory_of_thm th) = thy_name)
+ |> filter (fn (_, th) => Context.theory_name (Thm.theory_of_thm th) = thy_name)
fun do_while P f s list =
if P s then
@@ -52,7 +52,7 @@
if member (op =) S x then I
else insert (eq_set (op =)) (insert (op =) x S)) X) Ss []
fun check (s, th) =
- (case Logic.strip_horn (prop_of (Thm.unvarify_global th)) of
+ (case Logic.strip_horn (Thm.prop_of (Thm.unvarify_global th)) of
([], _) => (s, NONE)
| (ts, t) =>
let
--- a/src/HOL/Tools/Quotient/quotient_def.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_def.ML Wed Mar 04 19:53:18 2015 +0100
@@ -156,7 +156,7 @@
let
val lhs_eq =
thm
- |> prop_of
+ |> Thm.prop_of
|> Logic.dest_implies
|> fst
|> strip_all_body
@@ -178,7 +178,7 @@
val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs))
val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy
val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq
- val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq)
+ val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
fun after_qed thm_list lthy =
let
--- a/src/HOL/Tools/Quotient/quotient_info.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_info.ML Wed Mar 04 19:53:18 2015 +0100
@@ -125,7 +125,7 @@
Pretty.str "relation map:",
Pretty.str relmap,
Pretty.str "quot. theorem:",
- Syntax.pretty_term ctxt (prop_of quot_thm)])
+ Syntax.pretty_term ctxt (Thm.prop_of quot_thm)])
in
map prt_map (Symtab.dest (get_quotmaps (Context.Proof ctxt)))
|> Pretty.big_list "maps for type constructors:"
@@ -178,9 +178,9 @@
Pretty.str "relation:",
Syntax.pretty_term ctxt equiv_rel,
Pretty.str "equiv. thm:",
- Syntax.pretty_term ctxt (prop_of equiv_thm),
+ Syntax.pretty_term ctxt (Thm.prop_of equiv_thm),
Pretty.str "quot. thm:",
- Syntax.pretty_term ctxt (prop_of quot_thm)])
+ Syntax.pretty_term ctxt (Thm.prop_of quot_thm)])
in
map (prt_quot o snd) (Symtab.dest (get_quotients (Context.Proof ctxt)))
|> Pretty.big_list "quotients:"
@@ -218,7 +218,7 @@
Pretty.str ":=",
Syntax.pretty_term ctxt rconst,
Pretty.str "as",
- Syntax.pretty_term ctxt (prop_of def)])
+ Syntax.pretty_term ctxt (Thm.prop_of def)])
in
map prt_qconst (maps snd (Symtab.dest (get_quotconsts (Context.Proof ctxt))))
|> Pretty.big_list "quotient constants:"
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Wed Mar 04 19:53:18 2015 +0100
@@ -43,7 +43,7 @@
fun atomize_thm ctxt thm =
let
val thm' = Thm.legacy_freezeT (forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *)
- val thm'' = Object_Logic.atomize ctxt (cprop_of thm')
+ val thm'' = Object_Logic.atomize ctxt (Thm.cprop_of thm')
in
@{thm equal_elim_rule1} OF [thm'', thm']
end
@@ -73,10 +73,10 @@
fun prep_trm thy (x, (T, t)) =
- (cterm_of thy (Var (x, T)), cterm_of thy t)
+ (Thm.cterm_of thy (Var (x, T)), Thm.cterm_of thy t)
fun prep_ty thy (x, (S, ty)) =
- (ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
+ (Thm.ctyp_of thy (TVar (x, S)), Thm.ctyp_of thy ty)
fun get_match_inst thy pat trm =
let
@@ -100,8 +100,8 @@
let
val thy = Proof_Context.theory_of ctxt
fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm))
- val ty_inst = map (SOME o ctyp_of thy) [domain_type (fastype_of R2)]
- val trm_inst = map (SOME o cterm_of thy) [R2, R1]
+ val ty_inst = map (SOME o Thm.ctyp_of thy) [domain_type (fastype_of R2)]
+ val trm_inst = map (SOME o Thm.cterm_of thy) [R2, R1]
in
(case try (Drule.instantiate' ty_inst trm_inst) ball_bex_thm of
NONE => NONE
@@ -143,7 +143,7 @@
*)
fun reflp_get ctxt =
- map_filter (fn th => if prems_of th = [] then SOME (OF1 @{thm equivp_reflp} th) else NONE
+ map_filter (fn th => if Thm.prems_of th = [] then SOME (OF1 @{thm equivp_reflp} th) else NONE
handle THM _ => NONE) (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv}))
val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
@@ -194,22 +194,22 @@
end
fun quot_true_simple_conv ctxt fnctn ctrm =
- case term_of ctrm of
+ (case Thm.term_of ctrm of
(Const (@{const_name Quot_True}, _) $ x) =>
let
val fx = fnctn x;
val thy = Proof_Context.theory_of ctxt;
- val cx = cterm_of thy x;
- val cfx = cterm_of thy fx;
- val cxt = ctyp_of thy (fastype_of x);
- val cfxt = ctyp_of thy (fastype_of fx);
+ val cx = Thm.cterm_of thy x;
+ val cfx = Thm.cterm_of thy fx;
+ val cxt = Thm.ctyp_of thy (fastype_of x);
+ val cfxt = Thm.ctyp_of thy (fastype_of fx);
val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
in
Conv.rewr_conv thm ctrm
- end
+ end)
fun quot_true_conv ctxt fnctn ctrm =
- (case term_of ctrm of
+ (case Thm.term_of ctrm of
(Const (@{const_name Quot_True}, _) $ _) =>
quot_true_simple_conv ctxt fnctn ctrm
| _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm
@@ -238,8 +238,8 @@
val apply_rsp_tac =
Subgoal.FOCUS (fn {concl, asms, context,...} =>
let
- val bare_concl = HOLogic.dest_Trueprop (term_of concl)
- val qt_asm = find_qt_asm (map term_of asms)
+ val bare_concl = HOLogic.dest_Trueprop (Thm.term_of concl)
+ val qt_asm = find_qt_asm (map Thm.term_of asms)
in
case (bare_concl, qt_asm) of
(R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
@@ -251,8 +251,8 @@
val ty_b = fastype_of qt_arg
val ty_f = range_type (fastype_of f)
val thy = Proof_Context.theory_of context
- val ty_inst = map (SOME o (ctyp_of thy)) [ty_x, ty_b, ty_f]
- val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
+ val ty_inst = map (SOME o Thm.ctyp_of thy) [ty_x, ty_b, ty_f]
+ val t_inst = map (SOME o Thm.cterm_of thy) [R2, f, g, x, y];
val inst_thm = Drule.instantiate' ty_inst
([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3}
in
@@ -268,13 +268,13 @@
let
val thy = Proof_Context.theory_of ctxt
in
- case try (cterm_of thy) R of (* There can be loose bounds in R *)
+ case try (Thm.cterm_of thy) R of (* There can be loose bounds in R *)
SOME ctm =>
let
val ty = domain_type (fastype_of R)
in
- case try (Drule.instantiate' [SOME (ctyp_of thy ty)]
- [SOME (cterm_of thy R)]) @{thm equals_rsp} of
+ case try (Drule.instantiate' [SOME (Thm.ctyp_of thy ty)]
+ [SOME (Thm.cterm_of thy R)]) @{thm equals_rsp} of
SOME thm => rtac thm THEN' quotient_tac ctxt
| NONE => K no_tac
end
@@ -288,9 +288,9 @@
(let
val thy = Proof_Context.theory_of ctxt;
val (ty_a, ty_b) = dest_funT (fastype_of abs);
- val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
+ val ty_inst = map (SOME o Thm.ctyp_of thy) [ty_a, ty_b];
in
- case try (map (SOME o (cterm_of thy))) [rel, abs, rep] of
+ case try (map (SOME o Thm.cterm_of thy)) [rel, abs, rep] of
SOME t_inst =>
(case try (Drule.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
SOME inst_thm => (rtac inst_thm THEN' quotient_tac ctxt) i
@@ -421,7 +421,7 @@
(* expands all map_funs, except in front of the (bound) variables listed in xs *)
fun map_fun_simple_conv xs ctrm =
- (case term_of ctrm of
+ (case Thm.term_of ctrm of
((Const (@{const_name "map_fun"}, _) $ _ $ _) $ h $ _) =>
if member (op=) xs h
then Conv.all_conv ctrm
@@ -429,11 +429,11 @@
| _ => Conv.all_conv ctrm)
fun map_fun_conv xs ctxt ctrm =
- (case term_of ctrm of
+ (case Thm.term_of ctrm of
_ $ _ =>
(Conv.comb_conv (map_fun_conv xs ctxt) then_conv
map_fun_simple_conv xs) ctrm
- | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv ((term_of x)::xs) ctxt) ctxt ctrm
+ | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv (Thm.term_of x :: xs) ctxt) ctxt ctrm
| _ => Conv.all_conv ctrm)
fun map_fun_tac ctxt = CONVERSION (map_fun_conv [] ctxt)
@@ -472,22 +472,23 @@
make_inst_id is used
*)
fun lambda_prs_simple_conv ctxt ctrm =
- (case term_of ctrm of
+ (case Thm.term_of ctrm of
(Const (@{const_name map_fun}, _) $ r1 $ a2) $ (Abs _) =>
let
val thy = Proof_Context.theory_of ctxt
val (ty_b, ty_a) = dest_funT (fastype_of r1)
val (ty_c, ty_d) = dest_funT (fastype_of a2)
- val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
- val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
+ val tyinst = map (SOME o Thm.ctyp_of thy) [ty_a, ty_b, ty_c, ty_d]
+ val tinst = [NONE, NONE, SOME (Thm.cterm_of thy r1), NONE, SOME (Thm.cterm_of thy a2)]
val thm1 = Drule.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]}
val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1)
val thm3 = rewrite_rule ctxt @{thms id_apply[THEN eq_reflection]} thm2
val (insp, inst) =
if ty_c = ty_d
- then make_inst_id (term_of (Thm.lhs_of thm3)) (term_of ctrm)
- else make_inst (term_of (Thm.lhs_of thm3)) (term_of ctrm)
- val thm4 = Drule.instantiate_normalize ([], [(cterm_of thy insp, cterm_of thy inst)]) thm3
+ then make_inst_id (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm)
+ else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm)
+ val thm4 =
+ Drule.instantiate_normalize ([], [(Thm.cterm_of thy insp, Thm.cterm_of thy inst)]) thm3
in
Conv.rewr_conv thm4 ctrm
end
@@ -540,7 +541,7 @@
(* Tactic for Generalising Free Variables in a Goal *)
fun inst_spec ctrm =
- Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
+ Drule.instantiate' [SOME (Thm.ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
fun inst_spec_tac ctrms =
EVERY' (map (dtac o inst_spec) ctrms)
@@ -556,7 +557,7 @@
let
val thy = Proof_Context.theory_of ctxt
val vrs = Term.add_frees concl []
- val cvrs = map (cterm_of thy o Free) vrs
+ val cvrs = map (Thm.cterm_of thy o Free) vrs
val concl' = apply_under_Trueprop (all_list vrs) concl
val goal = Logic.mk_implies (concl', concl)
val rule = Goal.prove ctxt [] [] goal
@@ -616,10 +617,10 @@
handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
in
Drule.instantiate' []
- [SOME (cterm_of thy rtrm'),
- SOME (cterm_of thy reg_goal),
+ [SOME (Thm.cterm_of thy rtrm'),
+ SOME (Thm.cterm_of thy reg_goal),
NONE,
- SOME (cterm_of thy inj_goal)] procedure_thm
+ SOME (Thm.cterm_of thy inj_goal)] procedure_thm
end
@@ -676,9 +677,9 @@
handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
in
Drule.instantiate' []
- [SOME (cterm_of thy reg_goal),
+ [SOME (Thm.cterm_of thy reg_goal),
NONE,
- SOME (cterm_of thy inj_goal)] partiality_procedure_thm
+ SOME (Thm.cterm_of thy inj_goal)] partiality_procedure_thm
end
@@ -734,7 +735,7 @@
|> Thm.forall_intr_frees
|> atomize_thm ctxt
- val rule = procedure_inst ctxt (prop_of rthm') goal
+ val rule = procedure_inst ctxt (Thm.prop_of rthm') goal
in
(rtac rule THEN' rtac rthm') i
end)
@@ -757,7 +758,7 @@
fun lifted ctxt qtys simps rthm =
let
val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt
- val goal = Quotient_Term.derive_qtrm ctxt' qtys (prop_of rthm')
+ val goal = Quotient_Term.derive_qtrm ctxt' qtys (Thm.prop_of rthm')
in
Goal.prove ctxt' [] [] goal
(K (HEADGOAL (lift_single_tac ctxt' simps rthm')))
--- a/src/HOL/Tools/Quotient/quotient_term.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_term.ML Wed Mar 04 19:53:18 2015 +0100
@@ -355,7 +355,7 @@
| NONE => raise CODE_GEN ("get_relmap (no relation map function found for type " ^ s ^ ")"))
end
-fun is_id_quot thm = (prop_of thm = prop_of @{thm identity_quotient3})
+fun is_id_quot thm = (Thm.prop_of thm = Thm.prop_of @{thm identity_quotient3}) (* FIXME equality *)
open Lifting_Util
@@ -365,7 +365,7 @@
fun get_rel_from_quot_thm quot_thm =
let
- val (_ $ rel $ _ $ _) = (HOLogic.dest_Trueprop o prop_of) quot_thm
+ val (_ $ rel $ _ $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_thm
in
rel
end
--- a/src/HOL/Tools/Quotient/quotient_type.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Quotient/quotient_type.ML Wed Mar 04 19:53:18 2015 +0100
@@ -94,7 +94,7 @@
val (rty, qty) = (dest_funT o fastype_of) abs_fun
val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0)
- val Abs_body = (case (HOLogic.dest_Trueprop o prop_of) equiv_thm of
+ val Abs_body = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of
Const (@{const_name equivp}, _) $ _ => abs_fun_graph
| Const (@{const_name part_equivp}, _) $ rel =>
HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph)
@@ -112,20 +112,19 @@
fun setup_lifting_package quot3_thm equiv_thm opt_par_thm lthy =
let
- val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o prop_of) quot3_thm
+ val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot3_thm
val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy
val (rty, qty) = (dest_funT o fastype_of) abs_fun
val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty
val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name
- val (reflp_thm, quot_thm) = (case (HOLogic.dest_Trueprop o prop_of) equiv_thm of
- Const (@{const_name equivp}, _) $ _ =>
- (SOME (equiv_thm RS @{thm equivp_reflp2}),
- [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp})
+ val (reflp_thm, quot_thm) =
+ (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of
+ Const (@{const_name equivp}, _) $ _ =>
+ (SOME (equiv_thm RS @{thm equivp_reflp2}),
+ [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp})
| Const (@{const_name part_equivp}, _) $ _ =>
- (NONE,
- [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient})
- | _ => error "unsupported equivalence theorem"
- )
+ (NONE, [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient})
+ | _ => error "unsupported equivalence theorem")
in
lthy'
|> Lifting_Setup.setup_by_quotient quot_thm reflp_thm opt_par_thm
@@ -134,7 +133,7 @@
fun init_quotient_infr quot_thm equiv_thm opt_par_thm lthy =
let
- val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o prop_of) quot_thm
+ val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_thm
val (qtyp, rtyp) = (dest_funT o fastype_of) rep
val qty_full_name = (fst o dest_Type) qtyp
val quotients = {qtyp = qtyp, rtyp = rtyp, equiv_rel = rel, equiv_thm = equiv_thm,
--- a/src/HOL/Tools/SMT/verit_proof_parse.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/SMT/verit_proof_parse.ML Wed Mar 04 19:53:18 2015 +0100
@@ -37,7 +37,7 @@
fun step_of_assume j (_, th) =
VeriT_Proof.VeriT_Step {id = SMTLIB_Interface.assert_name_of_index (id_of_index j),
- rule = veriT_input_rule, prems = [], concl = prop_of th, fixes = []}
+ rule = veriT_input_rule, prems = [], concl = Thm.prop_of th, fixes = []}
val (actual_steps, _) = VeriT_Proof.parse typs terms output ctxt
val used_assert_ids = fold add_used_asserts_in_step actual_steps []
@@ -65,8 +65,8 @@
val helper_ids' = filter (fn (i, _) => i >= helpers_i) used_assms
val fact_helper_ts =
- map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, prop_of th)) helper_ids' @
- map (fn (_, ((s, _), th)) => (s, prop_of th)) fact_ids'
+ map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @
+ map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids'
val fact_helper_ids' =
map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids'
in
--- a/src/HOL/Tools/SMT/z3_replay.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/SMT/z3_replay.ML Wed Mar 04 19:53:18 2015 +0100
@@ -201,8 +201,8 @@
val helper_ids' = map_filter (try (fn (~1, idth) => idth)) iidths
val fact_helper_ts =
- map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, prop_of th)) helper_ids' @
- map (fn (_, ((s, _), th)) => (s, prop_of th)) fact_ids'
+ map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, Thm.prop_of th)) helper_ids' @
+ map (fn (_, ((s, _), th)) => (s, Thm.prop_of th)) fact_ids'
val fact_helper_ids' =
map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids' @ map (apsnd (fst o fst)) fact_ids'
in
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -75,7 +75,7 @@
val {context = ctxt, facts = chained, goal} = Proof.goal state
val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt
- val goal_t = Logic.list_implies (map prop_of chained @ hyp_ts, concl_t)
+ val goal_t = Logic.list_implies (map Thm.prop_of chained @ hyp_ts, concl_t)
fun try_methss [] [] = (used_facts, (preferred_meth, Play_Timed_Out Time.zeroTime))
| try_methss ress [] =
@@ -305,7 +305,7 @@
val problem =
{comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
factss = factss}
- val learn = mash_learn_proof ctxt params (prop_of goal) all_facts
+ val learn = mash_learn_proof ctxt params (Thm.prop_of goal) all_facts
val launch = launch_prover params mode output_result only learn
in
if mode = Auto_Try then
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML Wed Mar 04 19:53:18 2015 +0100
@@ -91,7 +91,7 @@
| is_rec_def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) = is_rec_eq t1 t2
| is_rec_def _ = false
-fun is_assum assms th = exists (fn ct => prop_of th aconv term_of ct) assms
+fun is_assum assms th = exists (fn ct => Thm.prop_of th aconv Thm.term_of ct) assms
fun is_chained chained = member Thm.eq_thm_prop chained
fun scope_of_thm global assms chained th =
@@ -131,7 +131,7 @@
if Termtab.is_empty css then
General
else
- let val t = prop_of th in
+ let val t = Thm.prop_of th in
(* FIXME: use structured name *)
if String.isSubstring ".induct" name andalso may_be_induction t then
Induction
@@ -224,7 +224,7 @@
val skolem_thesis = Name.skolem Auto_Bind.thesisN
fun is_that_fact th =
- exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (prop_of th)
+ exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (Thm.prop_of th)
andalso String.isSuffix sep_that (Thm.get_name_hint th)
datatype interest = Deal_Breaker | Interesting | Boring
@@ -266,7 +266,7 @@
combine_interests (interest_of_bool t) (interest_of_bool u)
| interest_of_prop _ _ = Deal_Breaker
- val t = prop_of th
+ val t = Thm.prop_of th
in
(interest_of_prop [] t <> Interesting andalso not (Thm.eq_thm_prop (@{thm ext}, th))) orelse
is_that_fact th
@@ -298,7 +298,7 @@
|> fst
fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote
-fun backquote_thm ctxt = backquote_term ctxt o prop_of
+fun backquote_thm ctxt = backquote_term ctxt o Thm.prop_of
(* TODO: rewrite to use nets and/or to reuse existing data structures *)
fun clasimpset_rule_table_of ctxt =
@@ -307,7 +307,7 @@
Termtab.empty
else
let
- fun add stature th = Termtab.update (normalize_vars (prop_of th), stature)
+ fun add stature th = Termtab.update (normalize_vars (Thm.prop_of th), stature)
val {safeIs, (* safeEs, *) hazIs, (* hazEs, *) ...} = ctxt |> claset_of |> Classical.rep_cs
val intros = Item_Net.content safeIs @ Item_Net.content hazIs
@@ -320,7 +320,7 @@
|> filter (curry (op =) Spec_Rules.Equational o fst)
|> maps (snd o snd)
|> filter_out (member Thm.eq_thm_prop risky_defs)
- |> List.partition (is_rec_def o prop_of)
+ |> List.partition (is_rec_def o Thm.prop_of)
val spec_intros = specs
|> filter (member (op =) [Spec_Rules.Inductive, Spec_Rules.Co_Inductive] o fst)
|> maps (snd o snd)
@@ -359,7 +359,7 @@
fun build_name_tables name_of facts =
let
- fun cons_thm (_, th) = Termtab.cons_list (normalize_vars (normalize_eq (prop_of th)), th)
+ fun cons_thm (_, th) = Termtab.cons_list (normalize_vars (normalize_eq (Thm.prop_of th)), th)
fun add_plain canon alias =
Symtab.update (Thm.get_name_hint alias, name_of (if_thm_before canon alias))
fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases
@@ -373,13 +373,13 @@
fun fact_distinct eq facts =
fold (fn fact as (_, th) =>
- Net.insert_term_safe (eq o apply2 (normalize_eq o prop_of o snd))
- (normalize_eq (prop_of th), fact))
+ Net.insert_term_safe (eq o apply2 (normalize_eq o Thm.prop_of o snd))
+ (normalize_eq (Thm.prop_of th), fact))
facts Net.empty
|> Net.entries
fun struct_induct_rule_on th =
- (case Logic.strip_horn (prop_of th) of
+ (case Logic.strip_horn (Thm.prop_of th) of
(prems, @{const Trueprop} $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
if not (is_TVar ind_T) andalso length prems > 1 andalso
exists (exists_subterm (curry (op aconv) p)) prems andalso
@@ -495,7 +495,7 @@
(j - 1,
if not (member Thm.eq_thm_prop add_ths th) andalso
(is_likely_tautology_too_meta_or_too_technical th orelse
- is_too_complex (prop_of th)) then
+ is_too_complex (Thm.prop_of th)) then
accum
else
let
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Wed Mar 04 19:53:18 2015 +0100
@@ -783,7 +783,7 @@
| order => order)
fun crude_thm_ord p =
- (case crude_theory_ord (apply2 theory_of_thm p) of
+ (case crude_theory_ord (apply2 Thm.theory_of_thm p) of
EQUAL =>
(* The hack below is necessary because of odd dependencies that are not reflected in the theory
comparison. *)
@@ -795,7 +795,7 @@
end
| ord => ord)
-val thm_less_eq = Theory.subthy o apply2 theory_of_thm
+val thm_less_eq = Theory.subthy o apply2 Thm.theory_of_thm
fun thm_less p = thm_less_eq p andalso not (thm_less_eq (swap p))
val freezeT = Type.legacy_freeze_type
@@ -807,7 +807,7 @@
| freeze (Free (s, T)) = Free (s, freezeT T)
| freeze t = t
-fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
+fun goal_of_thm thy = Thm.prop_of #> freeze #> Thm.cterm_of thy #> Goal.init
fun run_prover_for_mash ctxt params prover goal_name facts goal =
let
@@ -1120,8 +1120,8 @@
fun maximal_wrt_access_graph _ [] = []
| maximal_wrt_access_graph access_G ((fact as (_, th)) :: facts) =
- let val thy = theory_of_thm th in
- fact :: filter_out (fn (_, th') => strict_subthy (theory_of_thm th', thy)) facts
+ let val thy = Thm.theory_of_thm th in
+ fact :: filter_out (fn (_, th') => strict_subthy (Thm.theory_of_thm th', thy)) facts
|> map (nickname_of_thm o snd)
|> maximal_wrt_graph access_G
end
@@ -1163,11 +1163,11 @@
val chained = filter (fn ((_, (scope, _)), _) => scope = Chained) facts
fun fact_has_right_theory (_, th) =
- thy_name = Context.theory_name (theory_of_thm th)
+ thy_name = Context.theory_name (Thm.theory_of_thm th)
fun chained_or_extra_features_of factor (((_, stature), th), weight) =
- [prop_of th]
- |> features_of ctxt (theory_of_thm th) stature
+ [Thm.prop_of th]
+ |> features_of ctxt (Thm.theory_of_thm th) stature
|> map (rpair (weight * factor))
val {access_G, xtabs = ((num_facts, fact_tab), (num_feats, feat_tab)), ffds, freqs, ...} =
@@ -1380,7 +1380,7 @@
(learns, (num_nontrivial, next_commit, _)) =
let
val name = nickname_of_thm th
- val feats = features_of ctxt (theory_of_thm th) stature [prop_of th]
+ val feats = features_of ctxt (Thm.theory_of_thm th) stature [Thm.prop_of th]
val deps = these (deps_of status th)
val num_nontrivial = num_nontrivial |> not (null deps) ? Integer.add 1
val learns = (name, parents, feats, deps) :: learns
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML Wed Mar 04 19:53:18 2015 +0100
@@ -220,7 +220,7 @@
t
fun theory_const_prop_of fudge th =
- theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
+ theory_constify fudge (Context.theory_name (Thm.theory_of_thm th)) (Thm.prop_of th)
fun pair_consts_fact thy fudge fact =
(case fact |> snd |> theory_const_prop_of fudge |> pconsts_in_fact thy of
@@ -379,7 +379,7 @@
if Symtab.is_empty tab then
Symtab.empty
|> fold (add_pconsts_in_term thy) (map_filter (fn ((_, (sc', _)), th) =>
- if sc' = sc then SOME (prop_of th) else NONE) facts)
+ if sc' = sc then SOME (Thm.prop_of th) else NONE) facts)
else
tab
@@ -398,7 +398,7 @@
| SOME n => if n = length args then SOME tab else NONE))
| _ => SOME tab)
in
- aux (prop_of th) []
+ aux (Thm.prop_of th) []
end
(* FIXME: This is currently only useful for polymorphic type encodings. *)
@@ -421,7 +421,7 @@
val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
val add_pconsts = add_pconsts_in_term thy
val chained_ts =
- facts |> map_filter (try (fn ((_, (Chained, _)), th) => prop_of th))
+ facts |> map_filter (try (fn ((_, (Chained, _)), th) => Thm.prop_of th))
val chained_const_tab = Symtab.empty |> fold add_pconsts chained_ts
val goal_const_tab =
Symtab.empty
@@ -502,7 +502,7 @@
fold_aterms (curry (fn (Const (s', _), false) => s' = s | (_, b) => b)) t
false
fun uses_const_anywhere accepts s =
- exists (uses_const s o prop_of o snd) accepts orelse
+ exists (uses_const s o Thm.prop_of o snd) accepts orelse
exists (uses_const s) (concl_t :: hyp_ts)
fun add_set_const_thms accepts =
exists (uses_const_anywhere accepts) set_consts ? append set_thms
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML Wed Mar 04 19:53:18 2015 +0100
@@ -261,10 +261,10 @@
cache_value
else
facts
- |> not sound ? filter_out (is_dangerous_prop ctxt o prop_of o snd)
+ |> not sound ? filter_out (is_dangerous_prop ctxt o Thm.prop_of o snd)
|> take num_facts
|> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts
- |> map (apsnd prop_of)
+ |> map (apsnd Thm.prop_of)
|> (if waldmeister_new then
generate_waldmeister_problem ctxt hyp_ts concl_t
#> (fn (a,b,c,d,e) => (a,b,c,d,SOME e))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Wed Mar 04 19:53:18 2015 +0100
@@ -149,7 +149,7 @@
let
val (state, i, tool, message) = f ()
val ctxt = Proof.context_of state
- val goal = Logic.get_goal (prop_of (#goal (Proof.goal state))) i
+ val goal = Logic.get_goal (Thm.prop_of (#goal (Proof.goal state))) i
val hash = String.substring (SHA1.rep (SHA1.digest (hackish_string_of_term ctxt goal)), 0, 12)
in
File.append (Path.explode "$ISABELLE_HOME_USER/spy_sledgehammer")
--- a/src/HOL/Tools/TFL/dcterm.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/TFL/dcterm.ML Wed Mar 04 19:53:18 2015 +0100
@@ -70,7 +70,7 @@
val mk_hol_const = Thm.cterm_of @{theory HOL} o Const;
fun mk_exists (r as (Bvar, Body)) =
- let val ty = #T(rep_cterm Bvar)
+ let val ty = #T(Thm.rep_cterm Bvar)
val c = mk_hol_const(@{const_name Ex}, (ty --> HOLogic.boolT) --> HOLogic.boolT)
in capply c (uncurry cabs r) end;
@@ -88,12 +88,12 @@
* The primitives.
*---------------------------------------------------------------------------*)
fun dest_const ctm =
- (case #t(rep_cterm ctm)
+ (case #t(Thm.rep_cterm ctm)
of Const(s,ty) => {Name = s, Ty = ty}
| _ => raise ERR "dest_const" "not a constant");
fun dest_var ctm =
- (case #t(rep_cterm ctm)
+ (case #t(Thm.rep_cterm ctm)
of Var((s,i),ty) => {Name=s, Ty=ty}
| Free(s,ty) => {Name=s, Ty=ty}
| _ => raise ERR "dest_var" "not a variable");
--- a/src/HOL/Tools/TFL/post.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/TFL/post.ML Wed Mar 04 19:53:18 2015 +0100
@@ -26,7 +26,7 @@
*--------------------------------------------------------------------------*)
fun termination_goals rules =
map (Type.legacy_freeze o HOLogic.dest_Trueprop)
- (fold_rev (union (op aconv) o prems_of) rules []);
+ (fold_rev (union (op aconv) o Thm.prems_of) rules []);
(*---------------------------------------------------------------------------
* Three postprocessors are applied to the definition. It
@@ -62,17 +62,18 @@
handle Utils.ERR _ => false;
val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
-fun mk_meta_eq r = case concl_of r of
+fun mk_meta_eq r =
+ (case Thm.concl_of r of
Const(@{const_name Pure.eq},_)$_$_ => r
| _ $(Const(@{const_name HOL.eq},_)$_$_) => r RS eq_reflection
- | _ => r RS P_imp_P_eq_True
+ | _ => r RS P_imp_P_eq_True)
(*Is this the best way to invoke the simplifier??*)
fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L))
fun join_assums ctxt th =
let val thy = Thm.theory_of_thm th
- val tych = cterm_of thy
+ val tych = Thm.cterm_of thy
val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
val cntxtl = (#1 o USyntax.strip_imp) lhs (* cntxtl should = cntxtr *)
val cntxtr = (#1 o USyntax.strip_imp) rhs (* but union is solider *)
@@ -120,7 +121,7 @@
(*lcp: curry the predicate of the induction rule*)
fun curry_rule ctxt rl =
- Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl;
+ Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl;
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
fun meta_outer ctxt =
@@ -187,7 +188,7 @@
let val {functional,pats} = Prim.mk_functional thy eqs
val (thy, def) = Prim.wfrec_definition0 thy fid R functional
val ctxt = Proof_Context.transfer thy ctxt
- val (lhs, _) = Logic.dest_equals (prop_of def)
+ val (lhs, _) = Logic.dest_equals (Thm.prop_of def)
val {induct, rules, tcs} = simplify_defn strict thy ctxt congs wfs fid pats def
val rules' =
if strict then derive_init_eqs ctxt rules eqs
--- a/src/HOL/Tools/TFL/rules.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/TFL/rules.ML Wed Mar 04 19:53:18 2015 +0100
@@ -245,7 +245,8 @@
fun DISJ_CASESL disjth thl =
let val c = cconcl disjth
- fun eq th atm = exists (fn t => HOLogic.dest_Trueprop t aconv term_of atm) (Thm.hyps_of th)
+ fun eq th atm =
+ exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th)
val tml = Dcterm.strip_disj c
fun DL th [] = raise RULES_ERR "DISJ_CASESL" "no cases"
| DL th [th1] = PROVE_HYP th th1
@@ -264,8 +265,8 @@
val thy = Thm.theory_of_thm spec
val prop = Thm.prop_of spec
val x = hd (tl (Misc_Legacy.term_vars prop))
- val cTV = ctyp_of thy (type_of x)
- val gspec = Thm.forall_intr (cterm_of thy x) spec
+ val cTV = Thm.ctyp_of thy (type_of x)
+ val gspec = Thm.forall_intr (Thm.cterm_of thy x) spec
in
fun SPEC tm thm =
let val gspec' = Drule.instantiate_normalize ([(cTV, Thm.ctyp_of_term tm)], []) gspec
@@ -281,10 +282,10 @@
local val thy = Thm.theory_of_thm allI
val prop = Thm.prop_of allI
val [P] = Misc_Legacy.add_term_vars (prop, [])
- fun cty_theta s = map (fn (i, (S, ty)) => (ctyp_of s (TVar (i, S)), ctyp_of s ty))
+ fun cty_theta s = map (fn (i, (S, ty)) => (Thm.ctyp_of s (TVar (i, S)), Thm.ctyp_of s ty))
fun ctm_theta s = map (fn (i, (_, tm2)) =>
- let val ctm2 = cterm_of s tm2
- in (cterm_of s (Var(i,#T(rep_cterm ctm2))), ctm2)
+ let val ctm2 = Thm.cterm_of s tm2
+ in (Thm.cterm_of s (Var(i,#T(Thm.rep_cterm ctm2))), ctm2)
end)
fun certify s (ty_theta,tm_theta) =
(cty_theta s (Vartab.dest ty_theta),
@@ -300,7 +301,7 @@
val thm = Thm.implies_elim allI2 gth
val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm
val prop' = tp $ (A $ Abs(x,ty,M))
- in ALPHA thm (cterm_of thy prop')
+ in ALPHA thm (Thm.cterm_of thy prop')
end
end;
@@ -309,7 +310,7 @@
fun GEN_ALL thm =
let val thy = Thm.theory_of_thm thm
val prop = Thm.prop_of thm
- val tycheck = cterm_of thy
+ val tycheck = Thm.cterm_of thy
val vlist = map tycheck (Misc_Legacy.add_term_vars (prop, []))
in GENL vlist thm
end;
@@ -355,8 +356,8 @@
fun EXISTS (template,witness) thm =
let val thy = Thm.theory_of_thm thm
val prop = Thm.prop_of thm
- val P' = cterm_of thy P
- val x' = cterm_of thy x
+ val P' = Thm.cterm_of thy P
+ val x' = Thm.cterm_of thy x
val abstr = #2 (Dcterm.dest_comb template)
in
thm RS (cterm_instantiate[(P',abstr), (x',witness)] exI)
@@ -388,9 +389,9 @@
fun IT_EXISTS blist th =
let val thy = Thm.theory_of_thm th
- val tych = cterm_of thy
+ val tych = Thm.cterm_of thy
val blist' = map (apply2 Thm.term_of) blist
- fun ex v M = cterm_of thy (USyntax.mk_exists{Bvar=v,Body = M})
+ fun ex v M = Thm.cterm_of thy (USyntax.mk_exists{Bvar=v,Body = M})
in
fold_rev (fn (b as (r1,r2)) => fn thm =>
@@ -510,7 +511,7 @@
(* Note: Thm.rename_params_rule counts from 1, not 0 *)
fun rename thm =
let val thy = Thm.theory_of_thm thm
- val tych = cterm_of thy
+ val tych = Thm.cterm_of thy
val ants = Logic.strip_imp_prems (Thm.prop_of thm)
val news = get (ants,1,[])
in
@@ -662,7 +663,7 @@
val dummy = say (Display.string_of_thm ctxt thm)
(* Unquantified eliminate *)
fun uq_eliminate (thm,imp,thy) =
- let val tych = cterm_of thy
+ let val tych = Thm.cterm_of thy
val dummy = print_cterm ctxt "To eliminate:" (tych imp)
val ants = map tych (Logic.strip_imp_prems imp)
val eq = Logic.strip_imp_concl imp
@@ -682,7 +683,7 @@
orelse error "assertion failed in CONTEXT_REWRITE_RULE"
val imp_body1 = subst_free (ListPair.zip (args, vstrl))
imp_body
- val tych = cterm_of thy
+ val tych = Thm.cterm_of thy
val ants1 = map tych (Logic.strip_imp_prems imp_body1)
val eq1 = Logic.strip_imp_concl imp_body1
val Q = get_lhs eq1
@@ -711,7 +712,7 @@
in if (pbeta_redex Q) (length vlist)
then pq_eliminate (thm,thy,vlist,imp_body,Q)
else
- let val tych = cterm_of thy
+ let val tych = Thm.cterm_of thy
val ants1 = map tych ants
val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt
val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm
@@ -762,13 +763,13 @@
val antl = case rcontext of [] => []
| _ => [USyntax.list_mk_conj(map cncl rcontext)]
val TC = genl(USyntax.list_mk_imp(antl, A))
- val dummy = print_cterm ctxt "func:" (cterm_of thy func)
- val dummy = print_cterm ctxt "TC:" (cterm_of thy (HOLogic.mk_Trueprop TC))
+ val dummy = print_cterm ctxt "func:" (Thm.cterm_of thy func)
+ val dummy = print_cterm ctxt "TC:" (Thm.cterm_of thy (HOLogic.mk_Trueprop TC))
val dummy = tc_list := (TC :: !tc_list)
val nestedp = is_some (USyntax.find_term is_func TC)
val dummy = if nestedp then say "nested" else say "not_nested"
val th' = if nestedp then raise RULES_ERR "solver" "nested function"
- else let val cTC = cterm_of thy
+ else let val cTC = Thm.cterm_of thy
(HOLogic.mk_Trueprop TC)
in case rcontext of
[] => SPEC_ALL(ASSUME cTC)
@@ -781,8 +782,8 @@
in
(if (is_cong thm) then cong_prover else restrict_prover) ctxt thm
end
- val ctm = cprop_of th
- val names = Misc_Legacy.add_term_names (term_of ctm, [])
+ val ctm = Thm.cprop_of th
+ val names = Misc_Legacy.add_term_names (Thm.term_of ctm, [])
val th1 =
Raw_Simplifier.rewrite_cterm (false, true, false)
(prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm
--- a/src/HOL/Tools/Transfer/transfer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Transfer/transfer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -376,7 +376,7 @@
| _ $ _ $ _ => true
| _ => false
fun safe_transfer_rule_conv ctm =
- if is_transfer_rule (term_of ctm) then safe_Rel_conv ctm else Conv.all_conv ctm
+ if is_transfer_rule (Thm.term_of ctm) then safe_Rel_conv ctm else Conv.all_conv ctm
in
Conv.fconv_rule (Conv.prems_conv ~1 safe_transfer_rule_conv) thm
end
@@ -440,9 +440,9 @@
val thm0 = Thm.assume cprop
val (thm1, hyps) = zip ctxt' (thm0 :: thms) t u
val ((r1, x), y) = apfst Thm.dest_comb (Thm.dest_comb (Thm.dest_arg cprop))
- val r2 = Thm.dest_fun2 (Thm.dest_arg (cprop_of thm1))
- val (a1, (b1, _)) = apsnd dest_funcT (dest_funcT (ctyp_of_term r1))
- val (a2, (b2, _)) = apsnd dest_funcT (dest_funcT (ctyp_of_term r2))
+ val r2 = Thm.dest_fun2 (Thm.dest_arg (Thm.cprop_of thm1))
+ val (a1, (b1, _)) = apsnd dest_funcT (dest_funcT (Thm.ctyp_of_term r1))
+ val (a2, (b2, _)) = apsnd dest_funcT (dest_funcT (Thm.ctyp_of_term r2))
val tinsts = [SOME a1, SOME b1, SOME a2, SOME b2]
val insts = [SOME (Thm.dest_arg r1), SOME (Thm.dest_arg r2)]
val rule = Drule.instantiate' tinsts insts @{thm Rel_abs}
@@ -468,7 +468,7 @@
end
val r = mk_Rel (rel (fastype_of t) (fastype_of u))
val goal = HOLogic.mk_Trueprop (r $ t $ u)
- val rename = Thm.trivial (cterm_of thy goal)
+ val rename = Thm.trivial (Thm.cterm_of thy goal)
val (thm, hyps) = zip ctxt [] t u
in
Drule.implies_intr_list hyps (thm RS rename)
@@ -578,8 +578,9 @@
val relT = @{typ "bool => bool => bool"}
val idx = Thm.maxidx_of thm + 1
val thy = Proof_Context.theory_of ctxt
- fun tinst (a, _) = (ctyp_of thy (TVar ((a, idx), @{sort type})), cbool)
- fun inst (a, t) = (cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), cterm_of thy t)
+ fun tinst (a, _) = (Thm.ctyp_of thy (TVar ((a, idx), @{sort type})), cbool)
+ fun inst (a, t) =
+ (Thm.cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), Thm.cterm_of thy t)
in
thm
|> Thm.generalize (tfrees, rnames @ frees) idx
@@ -603,8 +604,9 @@
val relT = @{typ "bool => bool => bool"}
val idx = Thm.maxidx_of thm + 1
val thy = Proof_Context.theory_of ctxt
- fun tinst (a, _) = (ctyp_of thy (TVar ((a, idx), @{sort type})), cbool)
- fun inst (a, t) = (cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), cterm_of thy t)
+ fun tinst (a, _) = (Thm.ctyp_of thy (TVar ((a, idx), @{sort type})), cbool)
+ fun inst (a, t) =
+ (Thm.cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), Thm.cterm_of thy t)
in
thm
|> Thm.generalize (tfrees, rnames @ frees) idx
--- a/src/HOL/Tools/Transfer/transfer_bnf.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Wed Mar 04 19:53:18 2015 +0100
@@ -38,7 +38,7 @@
fun mk_pred pred_def args T =
let
- val pred_name = pred_def |> prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
+ val pred_name = pred_def |> Thm.prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
|> head_of |> fst o dest_Const
val argsT = map fastype_of args
in
@@ -91,7 +91,8 @@
fun generate_relation_constraint_goal ctxt bnf constraint_def =
let
- val constr_name = constraint_def |> prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
+ val constr_name =
+ constraint_def |> Thm.prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
|> head_of |> fst o dest_Const
val live = live_of_bnf bnf
val (((As, Bs), Ds), ctxt) = ctxt
@@ -345,7 +346,7 @@
val args = map mk_eq_onp preds
val cTs = map (SOME o Proof_Context.ctyp_of lthy) (maps (replicate 2) As)
val cts = map (SOME o Proof_Context.cterm_of lthy) args
- fun get_rhs thm = thm |> concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
+ fun get_rhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
fun is_eqn thm = can get_rhs thm
fun rel2pred_massage thm =
let
--- a/src/HOL/Tools/choice_specification.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/choice_specification.ML Wed Mar 04 19:53:18 2015 +0100
@@ -17,8 +17,8 @@
fun mk_definitional [] arg = arg
| mk_definitional ((thname, cname, covld) :: cos) (thy, thm) =
- (case HOLogic.dest_Trueprop (concl_of thm) of
- Const (@{const_name Ex},_) $ P =>
+ (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
+ Const (@{const_name Ex}, _) $ P =>
let
val ctype = domain_type (type_of P)
val cname_full = Sign.intern_const thy cname
@@ -85,7 +85,7 @@
val rew_imps = alt_props |>
map (Object_Logic.atomize ctxt o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
val props' = rew_imps |>
- map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
+ map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of)
fun proc_single prop =
let
@@ -136,8 +136,8 @@
let
fun inst_all thy v thm =
let
- val cv = cterm_of thy v
- val cT = ctyp_of_term cv
+ val cv = Thm.cterm_of thy v
+ val cT = Thm.ctyp_of_term cv
val spec' = instantiate' [SOME cT] [NONE, SOME cv] spec
in thm RS spec' end
fun remove_alls frees thm = fold (inst_all (Thm.theory_of_thm thm)) frees thm
--- a/src/HOL/Tools/cnf.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/cnf.ML Wed Mar 04 19:53:18 2015 +0100
@@ -138,7 +138,7 @@
(* adding new premises, then continues with the (i+1)-th premise *)
(* int -> Thm.thm -> Thm.thm *)
fun not_disj_to_prem i thm =
- if i > nprems_of thm then
+ if i > Thm.nprems_of thm then
thm
else
not_disj_to_prem (i+1)
@@ -165,7 +165,7 @@
(* ------------------------------------------------------------------------- *)
fun inst_thm thy ts thm =
- instantiate' [] (map (SOME o cterm_of thy) ts) thm;
+ instantiate' [] (map (SOME o Thm.cterm_of thy) ts) thm;
(* ------------------------------------------------------------------------- *)
(* Naive CNF transformation *)
@@ -261,12 +261,12 @@
let
val thy = Proof_Context.theory_of ctxt
fun conv ctxt ct =
- case term_of ct of
+ (case Thm.term_of ct of
Const _ $ Abs _ => Conv.comb_conv (conv ctxt) ct
| Abs _ => Conv.abs_conv (conv o snd) ctxt ct
| Const _ => Conv.all_conv ct
- | t => make t RS eq_reflection
- in conv ctxt (cterm_of thy t) RS meta_eq_to_obj_eq end
+ | t => make t RS eq_reflection)
+ in conv ctxt (Thm.cterm_of thy t) RS meta_eq_to_obj_eq end
fun make_nnf_thm_under_quantifiers ctxt =
make_under_quantifiers ctxt (make_nnf_thm (Proof_Context.theory_of ctxt))
@@ -293,14 +293,14 @@
fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
let
val thm1 = simp_True_False_thm thy x
- val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
+ val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
in
if x' = @{term False} then
simp_TF_conj_False_l OF [thm1] (* (x & y) = False *)
else
let
val thm2 = simp_True_False_thm thy y
- val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
+ val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
in
if x' = @{term True} then
simp_TF_conj_True_l OF [thm1, thm2] (* (x & y) = y' *)
@@ -315,14 +315,14 @@
| simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
let
val thm1 = simp_True_False_thm thy x
- val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
+ val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
in
if x' = @{term True} then
simp_TF_disj_True_l OF [thm1] (* (x | y) = True *)
else
let
val thm2 = simp_True_False_thm thy y
- val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
+ val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
in
if x' = @{term False} then
simp_TF_disj_False_l OF [thm1, thm2] (* (x | y) = y' *)
@@ -373,8 +373,8 @@
inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl (* (x' | y') = (x' | y') *)
val thm1 = make_cnf_thm_from_nnf x
val thm2 = make_cnf_thm_from_nnf y
- val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
- val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
+ val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
+ val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
val disj_thm = disj_cong OF [thm1, thm2] (* (x | y) = (x' | y') *)
in
iff_trans OF [disj_thm, make_cnf_disj_thm x' y']
@@ -385,10 +385,10 @@
(*###
val nnf_thm = make_nnf_thm thy t
*)
- val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) nnf_thm
+ val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
(* then simplify wrt. True/False (this should preserve NNF) *)
val simp_thm = simp_True_False_thm thy nnf
- val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) simp_thm
+ val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
(* finally, convert to CNF (this should preserve the simplification) *)
val cnf_thm = make_under_quantifiers ctxt make_cnf_thm_from_nnf simp
(* ###
@@ -451,7 +451,7 @@
val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_l (* ((Ex x') | y') = (Ex (x' | y')) *)
val var = new_free ()
val thm2 = make_cnfx_disj_thm (betapply (x', var)) y' (* (x' | y') = body' *)
- val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
+ val thm3 = Thm.forall_intr (Thm.cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
in
@@ -462,7 +462,7 @@
val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_r (* (x' | (Ex y')) = (Ex (x' | y')) *)
val var = new_free ()
val thm2 = make_cnfx_disj_thm x' (betapply (y', var)) (* (x' | y') = body' *)
- val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
+ val thm3 = Thm.forall_intr (Thm.cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x' | y') = body' *)
val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
in
@@ -472,8 +472,8 @@
inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl (* (x' | y') = (x' | y') *)
val thm1 = make_cnfx_thm_from_nnf x
val thm2 = make_cnfx_thm_from_nnf y
- val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
- val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
+ val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
+ val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
val disj_thm = disj_cong OF [thm1, thm2] (* (x | y) = (x' | y') *)
in
iff_trans OF [disj_thm, make_cnfx_disj_thm x' y']
@@ -484,7 +484,7 @@
val var = new_free ()
val body = HOLogic.mk_conj (HOLogic.mk_disj (x, var), HOLogic.mk_disj (y, HOLogic.Not $ var))
val thm2 = make_cnfx_thm_from_nnf body (* (x | v) & (y | ~v) = body' *)
- val thm3 = Thm.forall_intr (cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
+ val thm3 = Thm.forall_intr (Thm.cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
val thm4 = Thm.strip_shyps (thm3 COMP allI) (* ALL v. (x | v) & (y | ~v) = body' *)
val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x | v) & (y | ~v)) = (EX v. body') *)
in
@@ -496,10 +496,10 @@
(* ###
val nnf_thm = make_nnf_thm thy t
*)
- val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) nnf_thm
+ val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
(* then simplify wrt. True/False (this should preserve NNF) *)
val simp_thm = simp_True_False_thm thy nnf
- val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) simp_thm
+ val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
(* initialize var_id, in case the term already contains variables of the form "cnfx_<int>" *)
val _ = (var_id := fold (fn free => fn max =>
let
@@ -550,7 +550,7 @@
(* remove the original premises *)
THEN SELECT_GOAL (fn thm =>
let
- val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o prop_of) thm)
+ val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
in
PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
end) i;
@@ -564,7 +564,7 @@
(* cut the CNF formulas as new premises *)
Subgoal.FOCUS (fn {prems, ...} =>
let
- val cnfx_thms = map (make_cnfx_thm ctxt o HOLogic.dest_Trueprop o prop_of) prems
+ val cnfx_thms = map (make_cnfx_thm ctxt o HOLogic.dest_Trueprop o Thm.prop_of) prems
val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnfx_thms ~~ prems)
in
cut_facts_tac cut_thms 1
@@ -572,7 +572,7 @@
(* remove the original premises *)
THEN SELECT_GOAL (fn thm =>
let
- val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o prop_of) thm)
+ val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
in
PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
end) i;
--- a/src/HOL/Tools/coinduction.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/coinduction.ML Wed Mar 04 19:53:18 2015 +0100
@@ -24,15 +24,15 @@
fun ALLGOALS_SKIP skip tac st =
let fun doall n = if n = skip then all_tac else tac n THEN doall (n - 1)
- in doall (nprems_of st) st end;
+ in doall (Thm.nprems_of st) st end;
fun THEN_ALL_NEW_SKIP skip tac1 tac2 i st =
st |> (tac1 i THEN (fn st' =>
- Seq.INTERVAL tac2 (i + skip) (i + nprems_of st' - nprems_of st) st'));
+ Seq.INTERVAL tac2 (i + skip) (i + Thm.nprems_of st' - Thm.nprems_of st) st'));
fun DELETE_PREMS_AFTER skip tac i st =
let
- val n = nth (prems_of st) (i - 1) |> Logic.strip_assums_hyp |> length;
+ val n = nth (Thm.prems_of st) (i - 1) |> Logic.strip_assums_hyp |> length;
in
(THEN_ALL_NEW_SKIP skip tac (REPEAT_DETERM_N n o eresolve0_tac [thin_rl])) i st
end;
@@ -56,14 +56,14 @@
Object_Logic.atomize_prems_tac ctxt THEN'
DELETE_PREMS_AFTER skip (Subgoal.FOCUS (fn {concl, context = ctxt, params, prems, ...} =>
let
- val vars = raw_vars @ map (term_of o snd) params;
+ val vars = raw_vars @ map (Thm.term_of o snd) params;
val names_ctxt = ctxt
|> fold Variable.declare_names vars
|> fold Variable.declare_thm (raw_thm :: prems);
val thm_concl = Thm.cprop_of raw_thm |> strip_imp_concl;
val (rhoTs, rhots) = Thm.match (thm_concl, concl)
- |>> map (apply2 typ_of)
- ||> map (apply2 term_of);
+ |>> map (apply2 Thm.typ_of)
+ ||> map (apply2 Thm.term_of);
val xs = hd (Thm.prems_of raw_thm) |> HOLogic.dest_Trueprop |> strip_comb |> snd
|> map (subst_atomic_types rhoTs);
val raw_eqs = map (fn x => (x, AList.lookup op aconv rhots x |> the)) xs;
@@ -73,7 +73,7 @@
@{map 3} (fn name => fn T => fn (_, rhs) =>
HOLogic.mk_eq (Free (name, T), rhs))
names Ts raw_eqs;
- val phi = eqs @ map (HOLogic.dest_Trueprop o prop_of) prems
+ val phi = eqs @ map (HOLogic.dest_Trueprop o Thm.prop_of) prems
|> try (Library.foldr1 HOLogic.mk_conj)
|> the_default @{term True}
|> Ctr_Sugar_Util.list_exists_free vars
@@ -104,8 +104,9 @@
val inv_thms = init @ [last];
val eqs = take e inv_thms;
fun is_local_var t =
- member (fn (t, (_, t')) => t aconv (term_of t')) params t;
- val (eqs, assms') = filter_in_out (is_local_var o lhs_of_eq o prop_of) eqs;
+ member (fn (t, (_, t')) => t aconv (Thm.term_of t')) params t;
+ val (eqs, assms') =
+ filter_in_out (is_local_var o lhs_of_eq o Thm.prop_of) eqs;
val assms = assms' @ drop e inv_thms
in
HEADGOAL (Method.insert_tac (assms @ case_prems)) THEN
@@ -114,7 +115,7 @@
end) ctxt) THEN'
K (prune_params_tac ctxt))
#> Seq.maps (fn st =>
- CASES (Rule_Cases.make_common (Proof_Context.theory_of ctxt, prop_of st) cases) all_tac st)
+ CASES (Rule_Cases.make_common (Proof_Context.theory_of ctxt, Thm.prop_of st) cases) all_tac st)
end));
local
--- a/src/HOL/Tools/datatype_realizer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/datatype_realizer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -27,7 +27,7 @@
fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Old_Datatype_Aux.info) is thy =
let
val ctxt = Proof_Context.init_global thy;
- val cert = cterm_of thy;
+ val cert = Thm.cterm_of thy;
val recTs = Old_Datatype_Aux.get_rec_types descr;
val pnames =
@@ -107,7 +107,7 @@
make_pred i U T (mk_proj i is r) (Free (tname, T)))
(descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
val inst = map (apply2 cert) (map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induct))) ~~ ps);
+ (HOLogic.dest_Trueprop (Thm.concl_of induct))) ~~ ps);
val thm =
Goal.prove_internal ctxt (map cert prems) (cert concl)
@@ -144,7 +144,7 @@
(Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
end
| _ => AbsP ("H", SOME p, prf)))
- (rec_fns ~~ prems_of thm)
+ (rec_fns ~~ Thm.prems_of thm)
(Proofterm.proof_combP
(Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
@@ -160,7 +160,7 @@
fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Old_Datatype_Aux.info) thy =
let
val ctxt = Proof_Context.init_global thy;
- val cert = cterm_of thy;
+ val cert = Thm.cterm_of thy;
val rT = TFree ("'P", @{sort type});
val rT' = TVar (("'P", 0), @{sort type});
@@ -214,7 +214,7 @@
(Reconstruct.proof_of thm', map PBound (length prems - 1 downto 0))));
val prf' =
Extraction.abs_corr_shyps thy' exhaust []
- (map Var (Term.add_vars (prop_of exhaust) [])) (Reconstruct.proof_of exhaust);
+ (map Var (Term.add_vars (Thm.prop_of exhaust) [])) (Reconstruct.proof_of exhaust);
val r' =
Logic.varify_global (Abs ("y", T,
(fold_rev (Term.abs o dest_Free) rs
--- a/src/HOL/Tools/groebner.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/groebner.ML Wed Mar 04 19:53:18 2015 +0100
@@ -25,7 +25,7 @@
fun is_binop ct ct' =
(case Thm.term_of ct' of
- c $ _ $ _ => term_of ct aconv c
+ c $ _ $ _ => Thm.term_of ct aconv c
| _ => false);
fun dest_binary ct ct' =
@@ -351,7 +351,7 @@
(* We return an ideal_conv and the actual ring prover. *)
fun refute_disj rfn tm =
- case term_of tm of
+ case Thm.term_of tm of
Const(@{const_name HOL.disj},_)$_$_ =>
Drule.compose
(refute_disj rfn (Thm.dest_arg tm), 2,
@@ -362,11 +362,11 @@
fun mk_binop ct x y = Thm.apply (Thm.apply ct x) y
fun is_neg t =
- case term_of t of
+ case Thm.term_of t of
(Const(@{const_name Not},_)$_) => true
| _ => false;
fun is_eq t =
- case term_of t of
+ case Thm.term_of t of
(Const(@{const_name HOL.eq},_)$_$_) => true
| _ => false;
@@ -387,7 +387,7 @@
val strip_exists =
let fun h (acc, t) =
- case term_of t of
+ case Thm.term_of t of
Const (@{const_name Ex}, _) $ Abs _ =>
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
| _ => (acc,t)
@@ -395,7 +395,7 @@
end;
fun is_forall t =
- case term_of t of
+ case Thm.term_of t of
(Const(@{const_name All},_)$Abs(_,_,_)) => true
| _ => false;
@@ -435,12 +435,12 @@
fun sym_conv eq =
let val (l,r) = Thm.dest_binop eq
- in instantiate' [SOME (ctyp_of_term l)] [SOME l, SOME r] eq_commute
+ in instantiate' [SOME (Thm.ctyp_of_term l)] [SOME l, SOME r] eq_commute
end;
(* FIXME : copied from cqe.ML -- complex QE*)
fun conjuncts ct =
- case term_of ct of
+ case Thm.term_of ct of
@{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
| _ => [ct];
@@ -448,7 +448,7 @@
fun mk_conj_tab th =
let fun h acc th =
- case prop_of th of
+ case Thm.prop_of th of
@{term "Trueprop"}$(@{term HOL.conj}$_$_) =>
h (h acc (th RS conjunct2)) (th RS conjunct1)
| @{term "Trueprop"}$p => (p,th)::acc
@@ -459,7 +459,7 @@
fun prove_conj tab cjs =
case cjs of
- [c] => if is_conj (term_of c) then prove_conj tab (conjuncts c) else tab c
+ [c] => if is_conj (Thm.term_of c) then prove_conj tab (conjuncts c) else tab c
| c::cs => conjI OF [prove_conj tab [c], prove_conj tab cs];
fun conj_ac_rule eq =
@@ -467,8 +467,8 @@
val (l,r) = Thm.dest_equals eq
val ctabl = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} l))
val ctabr = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} r))
- fun tabl c = the (Termtab.lookup ctabl (term_of c))
- fun tabr c = the (Termtab.lookup ctabr (term_of c))
+ fun tabl c = the (Termtab.lookup ctabl (Thm.term_of c))
+ fun tabr c = the (Termtab.lookup ctabr (Thm.term_of c))
val thl = prove_conj tabl (conjuncts r) |> implies_intr_hyps
val thr = prove_conj tabr (conjuncts l) |> implies_intr_hyps
val eqI = instantiate' [] [SOME l, SOME r] @{thm iffI}
@@ -479,15 +479,15 @@
(* Conversion for the equivalence of existential statements where
EX quantifiers are rearranged differently *)
fun ext T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat Ex}
- fun mk_ex v t = Thm.apply (ext (ctyp_of_term v)) (Thm.lambda v t)
+ fun mk_ex v t = Thm.apply (ext (Thm.ctyp_of_term v)) (Thm.lambda v t)
-fun choose v th th' = case concl_of th of
+fun choose v th th' = case Thm.concl_of th of
@{term Trueprop} $ (Const(@{const_name Ex},_)$_) =>
let
- val p = (funpow 2 Thm.dest_arg o cprop_of) th
- val T = (hd o Thm.dest_ctyp o ctyp_of_term) p
+ val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th
+ val T = (hd o Thm.dest_ctyp o Thm.ctyp_of_term) p
val th0 = Conv.fconv_rule (Thm.beta_conversion true)
- (instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o cprop_of) th'] exE)
+ (instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE)
val pv = (Thm.rhs_of o Thm.beta_conversion true)
(Thm.apply @{cterm Trueprop} (Thm.apply p v))
val th1 = Thm.forall_intr v (Thm.implies_intr pv th')
@@ -504,7 +504,7 @@
val p = Thm.lambda v (Thm.dest_arg (Thm.cprop_of th))
in Thm.implies_elim
(Conv.fconv_rule (Thm.beta_conversion true)
- (instantiate' [SOME (ctyp_of_term v)] [SOME p, SOME v] @{thm exI}))
+ (instantiate' [SOME (Thm.ctyp_of_term v)] [SOME p, SOME v] @{thm exI}))
th
end
fun ex_eq_conv t =
@@ -515,19 +515,19 @@
val th = Thm.assume (Thm.apply @{cterm Trueprop} P)
val th1 = implies_intr_hyps (fold simple_choose vs' (fold mkexi vs th))
val th2 = implies_intr_hyps (fold simple_choose vs (fold mkexi vs' th))
- val p = (Thm.dest_arg o Thm.dest_arg1 o cprop_of) th1
- val q = (Thm.dest_arg o Thm.dest_arg o cprop_of) th1
+ val p = (Thm.dest_arg o Thm.dest_arg1 o Thm.cprop_of) th1
+ val q = (Thm.dest_arg o Thm.dest_arg o Thm.cprop_of) th1
in Thm.implies_elim (Thm.implies_elim (instantiate' [] [SOME p, SOME q] iffI) th1) th2
|> mk_meta_eq
end;
- fun getname v = case term_of v of
+ fun getname v = case Thm.term_of v of
Free(s,_) => s
| Var ((s,_),_) => s
| _ => "x"
fun mk_eq s t = Thm.apply (Thm.apply @{cterm "op == :: bool => _"} s) t
- fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v))
+ fun mk_exists v th = Drule.arg_cong_rule (ext (Thm.ctyp_of_term v))
(Thm.abstract_rule (getname v) v th)
fun simp_ex_conv ctxt =
Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)})
@@ -567,13 +567,13 @@
if length ideal = 2 then ideal else [eq_commute, eq_commute]
fun ring_dest_neg t =
let val (l,r) = Thm.dest_comb t
- in if Term.could_unify(term_of l,term_of ring_neg_tm) then r
+ in if Term.could_unify(Thm.term_of l, Thm.term_of ring_neg_tm) then r
else raise CTERM ("ring_dest_neg", [t])
end
fun field_dest_inv t =
let val (l,r) = Thm.dest_comb t in
- if Term.could_unify(term_of l, term_of field_inv_tm) then r
+ if Term.could_unify (Thm.term_of l, Thm.term_of field_inv_tm) then r
else raise CTERM ("field_dest_inv", [t])
end
val ring_dest_add = dest_binary ring_add_tm;
@@ -633,7 +633,7 @@
end)
handle CTERM _ =>
((let val (l,r) = ring_dest_pow tm
- in grob_pow vars (grobify_term vars l) ((term_of #> HOLogic.dest_number #> snd) r)
+ in grob_pow vars (grobify_term vars l) ((Thm.term_of #> HOLogic.dest_number #> snd) r)
end)
handle CTERM _ => error "grobify_term: unknown or invalid term")))))))));
val eq_tm = idom_thm |> concl |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_fun2;
@@ -649,7 +649,7 @@
val cjs = conjs tm
val rawvars =
fold_rev (fn eq => fn a => grobvars (Thm.dest_arg1 eq) (grobvars (Thm.dest_arg eq) a)) cjs []
- val vars = sort (fn (x, y) => Term_Ord.term_ord (term_of x, term_of y))
+ val vars = sort (fn (x, y) => Term_Ord.term_ord (Thm.term_of x, Thm.term_of y))
(distinct (op aconvc) rawvars)
in (vars,map (grobify_equation vars) cjs)
end;
@@ -739,7 +739,7 @@
let
fun mk_forall x p =
Thm.apply
- (Drule.cterm_rule (instantiate' [SOME (ctyp_of_term x)] [])
+ (Drule.cterm_rule (instantiate' [SOME (Thm.ctyp_of_term x)] [])
@{cpat "All:: (?'a => bool) => _"}) (Thm.lambda x p)
val avs = Thm.add_cterm_frees tm []
val P' = fold mk_forall avs tm
@@ -756,7 +756,7 @@
val th3 =
Thm.equal_elim
(Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [not_ex RS sym])
- (th2 |> cprop_of)) th2
+ (th2 |> Thm.cprop_of)) th2
in specl avs
([[[mk_object_eq th1, th3 RS PFalse'] MRS trans] MRS PFalse] MRS notnotD)
end
@@ -764,7 +764,7 @@
fun ideal tms tm ord =
let
val rawvars = fold_rev grobvars (tm::tms) []
- val vars = sort ord (distinct (fn (x,y) => (term_of x) aconv (term_of y)) rawvars)
+ val vars = sort ord (distinct (fn (x,y) => (Thm.term_of x) aconv (Thm.term_of y)) rawvars)
val pols = map (grobify_term vars) tms
val pol = grobify_term vars tm
val cert = grobner_ideal vars pols pol
@@ -866,7 +866,7 @@
val vars = filter (fn v => free_in v eq) evs
val (qs,p) = isolate_monomials vars eq
val rs = ideal (qs @ ps) p
- (fn (s,t) => Term_Ord.term_ord (term_of s, term_of t))
+ (fn (s,t) => Term_Ord.term_ord (Thm.term_of s, Thm.term_of t))
in (eq, take (length qs) rs ~~ vars)
end;
fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p));
@@ -888,7 +888,7 @@
fun find_term bounds tm =
- (case term_of tm of
+ (case Thm.term_of tm of
Const (@{const_name HOL.eq}, T) $ _ $ _ =>
if domain_type T = HOLogic.boolT then find_args bounds tm
else Thm.dest_arg tm
@@ -919,7 +919,7 @@
| SOME (res as (theory, {is_const = _, dest_const,
mk_const, conv = ring_eq_conv})) =>
SOME (ring_and_ideal_conv theory
- dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
+ dest_const (mk_const (Thm.ctyp_of_term tm)) (ring_eq_conv ctxt)
(Semiring_Normalizer.semiring_normalize_wrapper ctxt res)))
fun presimplify ctxt add_thms del_thms =
@@ -941,11 +941,11 @@
| THM _ => no_tac);
local
- fun lhs t = case term_of t of
+ fun lhs t = case Thm.term_of t of
Const(@{const_name HOL.eq},_)$_$_ => Thm.dest_arg1 t
| _=> raise CTERM ("ideal_tac - lhs",[t])
fun exitac NONE = no_tac
- | exitac (SOME y) = rtac (instantiate' [SOME (ctyp_of_term y)] [NONE,SOME y] exI) 1
+ | exitac (SOME y) = rtac (instantiate' [SOME (Thm.ctyp_of_term y)] [NONE,SOME y] exI) 1
val claset = claset_of @{context}
in
--- a/src/HOL/Tools/inductive.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/inductive.ML Wed Mar 04 19:53:18 2015 +0100
@@ -260,7 +260,7 @@
fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD})
handle THM _ => thm RS @{thm le_boolD}
in
- (case concl_of thm of
+ (case Thm.concl_of thm of
Const (@{const_name Pure.eq}, _) $ _ $ _ => eq_to_mono (thm RS meta_eq_to_obj_eq)
| _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => eq_to_mono thm
| _ $ (Const (@{const_name Orderings.less_eq}, _) $ _ $ _) =>
@@ -636,7 +636,7 @@
| _ => error
("equations matching pattern " ^ Syntax.string_of_term ctxt prop ^ " is not unique"));
val inst =
- map (fn v => (cterm_of thy (Var v), cterm_of thy (Envir.subst_term subst (Var v))))
+ map (fn v => (Thm.cterm_of thy (Var v), Thm.cterm_of thy (Envir.subst_term subst (Var v))))
(Term.add_vars (lhs_of eq) []);
in
Drule.cterm_instantiate inst eq
@@ -1101,12 +1101,12 @@
fun arities_of induct =
map (fn (_ $ t $ u) =>
(fst (dest_Const (head_of t)), length (snd (strip_comb u))))
- (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)));
(* read off parameters of inductive predicate from raw induction rule *)
fun params_of induct =
let
- val (_ $ t $ u :: _) = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct));
+ val (_ $ t $ u :: _) = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct));
val (_, ts) = strip_comb t;
val (_, us) = strip_comb u;
in
@@ -1114,7 +1114,7 @@
end;
val pname_of_intr =
- concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst;
+ Thm.concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst;
(* partition introduction rules according to predicate name *)
fun gen_partition_rules f induct intros =
@@ -1131,9 +1131,9 @@
(* infer order of variables in intro rules from order of quantifiers in elim rule *)
fun infer_intro_vars elim arity intros =
let
- val thy = theory_of_thm elim;
- val _ :: cases = prems_of elim;
- val used = map (fst o fst) (Term.add_vars (prop_of elim) []);
+ val thy = Thm.theory_of_thm elim;
+ val _ :: cases = Thm.prems_of elim;
+ val used = map (fst o fst) (Term.add_vars (Thm.prop_of elim) []);
fun mtch (t, u) =
let
val params = Logic.strip_params t;
@@ -1150,7 +1150,7 @@
map (Envir.subst_term tab) vars
end
in
- map (mtch o apsnd prop_of) (cases ~~ intros)
+ map (mtch o apsnd Thm.prop_of) (cases ~~ intros)
end;
--- a/src/HOL/Tools/inductive_realizer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -58,16 +58,16 @@
fun dt_of_intrs thy vs nparms intrs =
let
- val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []);
+ val iTs = rev (Term.add_tvars (Thm.prop_of (hd intrs)) []);
val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop
- (Logic.strip_imp_concl (prop_of (hd intrs))));
+ (Logic.strip_imp_concl (Thm.prop_of (hd intrs))));
val params = map dest_Var (take nparms ts);
val tname = Binding.name (space_implode "_" (Long_Name.base_name s ^ "T" :: vs));
fun constr_of_intr intr = (Binding.name (Long_Name.base_name (name_of_thm intr)),
- map (Logic.unvarifyT_global o snd) (subtract (op =) params (rev (Term.add_vars (prop_of intr) []))) @
- filter_out (equal Extraction.nullT) (map
- (Logic.unvarifyT_global o Extraction.etype_of thy vs []) (prems_of intr)),
- NoSyn);
+ map (Logic.unvarifyT_global o snd)
+ (subtract (op =) params (rev (Term.add_vars (Thm.prop_of intr) []))) @
+ filter_out (equal Extraction.nullT)
+ (map (Logic.unvarifyT_global o Extraction.etype_of thy vs []) (Thm.prems_of intr)), NoSyn);
in
((tname, map (rpair dummyS) (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs), NoSyn),
map constr_of_intr intrs)
@@ -95,7 +95,7 @@
fun mk_realizes_eqn n vs nparms intrs =
let
- val intr = map_types Type.strip_sorts (prop_of (hd intrs));
+ val intr = map_types Type.strip_sorts (Thm.prop_of (hd intrs));
val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl intr);
val iTs = rev (Term.add_tvars intr []);
val Tvs = map TVar iTs;
@@ -138,7 +138,7 @@
val ctxt = Proof_Context.init_global thy
val args = map (Free o apfst fst o dest_Var) ivs;
val args' = map (Free o apfst fst)
- (subtract (op =) params (Term.add_vars (prop_of intr) []));
+ (subtract (op =) params (Term.add_vars (Thm.prop_of intr) []));
val rule' = strip_all rule;
val conclT = Extraction.etype_of thy vs [] (Logic.strip_imp_concl rule');
val used = map (fst o dest_Free) args;
@@ -203,10 +203,10 @@
fun indrule_realizer thy induct raw_induct rsets params vs rec_names rss intrs dummies =
let
- val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct));
+ val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of raw_induct));
val premss = map_filter (fn (s, rs) => if member (op =) rsets s then
- SOME (rs, map (fn (_, r) => nth (prems_of raw_induct)
- (find_index (fn prp => prp = prop_of r) (map prop_of intrs))) rs) else NONE) rss;
+ SOME (rs, map (fn (_, r) => nth (Thm.prems_of raw_induct)
+ (find_index (fn prp => prp = Thm.prop_of r) (map Thm.prop_of intrs))) rs) else NONE) rss;
val fs = maps (fn ((intrs, prems), dummy) =>
let
val fs = map (fn (rule, (ivs, intr)) =>
@@ -257,13 +257,13 @@
fun mk_realizer thy vs (name, rule, rrule, rlz, rt) =
let
- val rvs = map fst (relevant_vars (prop_of rule));
- val xs = rev (Term.add_vars (prop_of rule) []);
+ val rvs = map fst (relevant_vars (Thm.prop_of rule));
+ val xs = rev (Term.add_vars (Thm.prop_of rule) []);
val vs1 = map Var (filter_out (fn ((a, _), _) => member (op =) rvs a) xs);
- val rlzvs = rev (Term.add_vars (prop_of rrule) []);
+ val rlzvs = rev (Term.add_vars (Thm.prop_of rrule) []);
val vs2 = map (fn (ixn, _) => Var (ixn, (the o AList.lookup (op =) rlzvs) ixn)) xs;
val rs = map Var (subtract (op = o apply2 fst) xs rlzvs);
- val rlz' = fold_rev Logic.all rs (prop_of rrule)
+ val rlz' = fold_rev Logic.all rs (Thm.prop_of rrule)
in (name, (vs,
if rt = Extraction.nullt then rt else fold_rev lambda vs1 rt,
Extraction.abs_corr_shyps thy rule vs vs2
@@ -277,7 +277,7 @@
let
val qualifier = Long_Name.qualifier (name_of_thm induct);
val inducts = Global_Theory.get_thms thy (Long_Name.qualify qualifier "inducts");
- val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []);
+ val iTs = rev (Term.add_tvars (Thm.prop_of (hd intrs)) []);
val ar = length vs + length iTs;
val params = Inductive.params_of raw_induct;
val arities = Inductive.arities_of raw_induct;
@@ -317,21 +317,21 @@
if null dt_names then []
else #rec_rewrites (Old_Datatype_Data.the_info thy2 (hd dt_names));
val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o
- HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) rec_thms);
+ HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) rec_thms);
val (constrss, _) = fold_map (fn (s, rs) => fn (recs, dummies) =>
if member (op =) rsets s then
let
val (d :: dummies') = dummies;
val (recs1, recs2) = chop (length rs) (if d then tl recs else recs)
in (map (head_of o hd o rev o snd o strip_comb o fst o
- HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1, (recs2, dummies'))
+ HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) recs1, (recs2, dummies'))
end
else (replicate (length rs) Extraction.nullt, (recs, dummies)))
rss (rec_thms, dummies);
val rintrs = map (fn (intr, c) => attach_typeS (Envir.eta_contract
(Extraction.realizes_of thy2 vs
(if c = Extraction.nullt then c else list_comb (c, map Var (rev
- (subtract (op =) params' (Term.add_vars (prop_of intr) []))))) (prop_of intr))))
+ (subtract (op =) params' (Term.add_vars (Thm.prop_of intr) []))))) (Thm.prop_of intr))))
(maps snd rss ~~ flat constrss);
val (rlzpreds, rlzpreds') =
rintrs |> map (fn rintr =>
@@ -365,17 +365,17 @@
val Ps = map_filter (fn _ $ M $ P => if member (op =) rsets (pred_of M) then
SOME (fst (fst (dest_Var (head_of P)))) else NONE)
- (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)));
+ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of raw_induct)));
fun add_ind_realizer Ps thy =
let
val vs' = rename (map (apply2 (fst o fst o dest_Var))
(params ~~ List.take (snd (strip_comb (HOLogic.dest_Trueprop
- (hd (prems_of (hd inducts))))), nparms))) vs;
+ (hd (Thm.prems_of (hd inducts))))), nparms))) vs;
val rs = indrule_realizer thy induct raw_induct rsets params'
(vs' @ Ps) rec_names rss' intrs dummies;
val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r
- (prop_of ind)) (rs ~~ inducts);
+ (Thm.prop_of ind)) (rs ~~ inducts);
val used = fold Term.add_free_names rlzs [];
val rnames = Name.variant_list used (replicate (length inducts) "r");
val rnames' = Name.variant_list
@@ -421,24 +421,24 @@
(** realizer for elimination rules **)
val case_names = map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o
- HOLogic.dest_Trueprop o prop_of o hd) case_thms;
+ HOLogic.dest_Trueprop o Thm.prop_of o hd) case_thms;
fun add_elim_realizer Ps
(((((elim, elimR), intrs), case_thms), case_name), dummy) thy =
let
- val (prem :: prems) = prems_of elim;
+ val (prem :: prems) = Thm.prems_of elim;
fun reorder1 (p, (_, intr)) =
fold (fn ((s, _), T) => Logic.all (Free (s, T)))
- (subtract (op =) params' (Term.add_vars (prop_of intr) []))
+ (subtract (op =) params' (Term.add_vars (Thm.prop_of intr) []))
(strip_all p);
fun reorder2 ((ivs, intr), i) =
- let val fs = subtract (op =) params' (Term.add_vars (prop_of intr) [])
+ let val fs = subtract (op =) params' (Term.add_vars (Thm.prop_of intr) [])
in fold (lambda o Var) fs (list_comb (Bound (i + length ivs), ivs)) end;
val p = Logic.list_implies
- (map reorder1 (prems ~~ intrs) @ [prem], concl_of elim);
+ (map reorder1 (prems ~~ intrs) @ [prem], Thm.concl_of elim);
val T' = Extraction.etype_of thy (vs @ Ps) [] p;
val T = if dummy then (HOLogic.unitT --> body_type T') --> T' else T';
- val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (prems_of elim);
+ val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (Thm.prems_of elim);
val r =
if null Ps then Extraction.nullt
else
@@ -449,7 +449,7 @@
else []) @
map reorder2 (intrs ~~ (length prems - 1 downto 0)) @
[Bound (length prems)]));
- val rlz = Extraction.realizes_of thy (vs @ Ps) r (prop_of elim);
+ val rlz = Extraction.realizes_of thy (vs @ Ps) r (Thm.prop_of elim);
val rlz' = attach_typeS (strip_all (Logic.unvarify_global rlz));
val rews = map mk_meta_eq case_thms;
val thm = Goal.prove_global thy []
@@ -474,7 +474,7 @@
val thy5 = Extraction.add_realizers_i
(map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) =>
(name_of_thm rule, rule, rrule, rlz,
- list_comb (c, map Var (subtract (op =) params' (rev (Term.add_vars (prop_of rule) []))))))
+ list_comb (c, map Var (subtract (op =) params' (rev (Term.add_vars (Thm.prop_of rule) []))))))
(maps snd rss ~~ #intrs ind_info ~~ rintrs ~~ flat constrss))) thy4;
val elimps = map_filter (fn ((s, intrs), p) =>
if member (op =) rsets s then SOME (p, intrs) else NONE)
@@ -482,7 +482,7 @@
val thy6 =
fold (fn p as (((((elim, _), _), _), _), _) =>
add_elim_realizer [] p #>
- add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (concl_of elim))))] p)
+ add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (Thm.concl_of elim))))] p)
(elimps ~~ case_thms ~~ case_names ~~ dummies) thy5;
in Sign.restore_naming thy thy6 end;
@@ -492,7 +492,7 @@
val (_, {intrs, induct, raw_induct, elims, ...}) =
Inductive.the_inductive (Proof_Context.init_global thy) name;
val vss = sort (int_ord o apply2 length)
- (subsets (map fst (relevant_vars (concl_of (hd intrs)))))
+ (subsets (map fst (relevant_vars (Thm.concl_of (hd intrs)))))
in
fold_rev (add_ind_realizer rsets intrs induct raw_induct elims) vss thy
end
@@ -501,8 +501,8 @@
let
fun err () = error "ind_realizer: bad rule";
val sets =
- (case HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of thm)) of
- [_] => [pred_of (HOLogic.dest_Trueprop (hd (prems_of thm)))]
+ (case HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of thm)) of
+ [_] => [pred_of (HOLogic.dest_Trueprop (hd (Thm.prems_of thm)))]
| xs => map (pred_of o fst o HOLogic.dest_imp) xs)
handle TERM _ => err () | List.Empty => err ();
in
--- a/src/HOL/Tools/inductive_set.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/inductive_set.ML Wed Mar 04 19:53:18 2015 +0100
@@ -44,7 +44,7 @@
val thy = Proof_Context.theory_of ctxt;
fun close p t f =
let val vs = Term.add_vars t []
- in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
+ in Drule.instantiate' [] (rev (map (SOME o Thm.cterm_of thy o Var) vs))
(p (fold (Logic.all o Var) vs t) f)
end;
fun mkop @{const_name HOL.conj} T x =
@@ -93,8 +93,8 @@
in
if forall is_none rews then NONE
else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1)
- (map2 (fn SOME r => K r | NONE => Thm.reflexive o cterm_of thy)
- rews ts) (Thm.reflexive (cterm_of thy h)))
+ (map2 (fn SOME r => K r | NONE => Thm.reflexive o Thm.cterm_of thy)
+ rews ts) (Thm.reflexive (Thm.cterm_of thy h)))
end
| NONE => NONE)
| _ => NONE
@@ -117,7 +117,7 @@
Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct =>
Thm.transitive (Thm.eta_conversion ct)
(Thm.symmetric (Thm.eta_conversion
- (cterm_of (theory_of_cterm ct) (eta_contract p (term_of ct)))))));
+ (Thm.cterm_of (Thm.theory_of_cterm ct) (eta_contract p (Thm.term_of ct)))))));
(***********************************************************)
@@ -208,8 +208,8 @@
val x' = map_type
(K (Ts @ HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x;
in
- (cterm_of thy x,
- cterm_of thy (fold_rev (Term.abs o pair "x") Ts
+ (Thm.cterm_of thy x,
+ Thm.cterm_of thy (fold_rev (Term.abs o pair "x") Ts
(HOLogic.Collect_const U $
HOLogic.mk_psplits ps U HOLogic.boolT
(list_comb (x', map Bound (length Ts - 1 downto 0))))))
@@ -217,7 +217,7 @@
fun mk_to_pred_eq ctxt p fs optfs' T thm =
let
- val thy = theory_of_thm thm;
+ val thy = Thm.theory_of_thm thm;
val insts = mk_to_pred_inst thy fs;
val thm' = Thm.instantiate ([], insts) thm;
val thm'' =
@@ -229,11 +229,11 @@
val Ts = HOLogic.strip_ptupleT fs' U;
(* FIXME: should cterm_instantiate increment indexes? *)
val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
- val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
+ val (arg_cong_f, _) = arg_cong' |> Thm.cprop_of |> Drule.strip_imp_concl |>
Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
in
thm' RS (Drule.cterm_instantiate [(arg_cong_f,
- cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT,
+ Thm.cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT,
HOLogic.Collect_const U $ HOLogic.mk_psplits fs' U
HOLogic.boolT (Bound 0))))] arg_cong' RS sym)
end)
@@ -249,7 +249,7 @@
exception Malformed of string;
fun add context thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
- (case prop_of thm of
+ (case Thm.prop_of thm of
Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) $ lhs $ rhs) =>
(case body_type T of
@{typ bool} =>
@@ -318,7 +318,8 @@
let val rules' = map mk_meta_eq rules
in
Simplifier.simproc_global_i @{theory HOL} "to_pred" [anyt]
- (fn ctxt => (lookup_rule (Proof_Context.theory_of ctxt) (prop_of #> Logic.dest_equals) rules'))
+ (fn ctxt =>
+ lookup_rule (Proof_Context.theory_of ctxt) (Thm.prop_of #> Logic.dest_equals) rules')
end;
fun to_pred_proc thy rules t = case lookup_rule thy I rules t of
@@ -334,7 +335,7 @@
val {to_pred_simps, set_arities, pred_arities, ...} =
fold (add context) thms (Data.get context);
val fs = filter (is_Var o fst)
- (infer_arities thy set_arities (NONE, prop_of thm) []);
+ (infer_arities thy set_arities (NONE, Thm.prop_of thm) []);
(* instantiate each set parameter with {(x, y). p x y} *)
val insts = mk_to_pred_inst thy fs
in
@@ -358,7 +359,7 @@
val {to_set_simps, pred_arities, ...} =
fold (add context) thms (Data.get context);
val fs = filter (is_Var o fst)
- (infer_arities thy pred_arities (NONE, prop_of thm) []);
+ (infer_arities thy pred_arities (NONE, Thm.prop_of thm) []);
(* instantiate each predicate parameter with %x y. (x, y) : s *)
val insts = map (fn (x, ps) =>
let
@@ -369,8 +370,8 @@
val T = HOLogic.mk_ptupleT ps Us;
val x' = map_type (K (Rs ---> HOLogic.mk_setT T)) x
in
- (cterm_of thy x,
- cterm_of thy (fold_rev (Term.abs o pair "x") Ts
+ (Thm.cterm_of thy x,
+ Thm.cterm_of thy (fold_rev (Term.abs o pair "x") Ts
(HOLogic.mk_mem (HOLogic.mk_ptuple ps T (map Bound (k downto 0)),
list_comb (x', map Bound (l - 1 downto k + 1))))))
end) fs;
@@ -397,7 +398,7 @@
if exists_Const (fn (s, _) => case Symtab.lookup set_arities s of
NONE => false
| SOME arities => exists (fn (_, (xs, _)) =>
- forall is_none xs) arities) (prop_of thm)
+ forall is_none xs) arities) (Thm.prop_of thm)
then
thm |>
Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimprocs
@@ -465,7 +466,7 @@
(list_comb (c', params1))))
end) |> split_list |>> split_list;
val eqns' = eqns @
- map (prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq)
+ map (Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq)
(mem_Collect_eq :: @{thm split_conv} :: to_pred_simps);
(* predicate version of the introduction rules *)
--- a/src/HOL/Tools/int_arith.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/int_arith.ML Wed Mar 04 19:53:18 2015 +0100
@@ -25,8 +25,8 @@
val lhss0 = [@{cpat "0::?'a::ring"}];
fun proc0 phi ctxt ct =
- let val T = ctyp_of_term ct
- in if typ_of T = @{typ int} then NONE else
+ let val T = Thm.ctyp_of_term ct
+ in if Thm.typ_of T = @{typ int} then NONE else
SOME (instantiate' [SOME T] [] zeroth)
end;
@@ -39,8 +39,8 @@
val lhss1 = [@{cpat "1::?'a::ring_1"}];
fun proc1 phi ctxt ct =
- let val T = ctyp_of_term ct
- in if typ_of T = @{typ int} then NONE else
+ let val T = Thm.ctyp_of_term ct
+ in if Thm.typ_of T = @{typ int} then NONE else
SOME (instantiate' [SOME T] [] oneth)
end;
@@ -67,7 +67,7 @@
addsimprocs [zero_to_of_int_zero_simproc,one_to_of_int_one_simproc]);
fun sproc phi ctxt ct =
- if check (term_of ct) then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)
+ if check (Thm.term_of ct) then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)
else NONE;
val lhss' =
--- a/src/HOL/Tools/legacy_transfer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/legacy_transfer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -21,7 +21,7 @@
(* data administration *)
-val direction_of = Thm.dest_binop o Thm.dest_arg o cprop_of;
+val direction_of = Thm.dest_binop o Thm.dest_arg o Thm.cprop_of;
val transfer_morphism_key = Drule.strip_imp_concl (Thm.cprop_of @{thm transfer_morphismI});
@@ -89,7 +89,7 @@
val tys = map snd (Term.add_vars t []);
val _ = if null tys then error "Transfer: unable to guess instance" else ();
val tyss = splits (curry Type.could_unify) tys;
- val get_ty = typ_of o ctyp_of_term o fst o direction_of;
+ val get_ty = Thm.typ_of o Thm.ctyp_of_term o fst o direction_of;
val insts = map_filter (fn tys => get_first (fn (k, e) =>
if Type.could_unify (hd tys, range_type (get_ty k))
then SOME (direction_of k, transfer_rules_of e)
@@ -114,8 +114,8 @@
(* determine variables to transfer *)
val ctxt3 = ctxt2
|> Variable.declare_thm thm
- |> Variable.declare_term (term_of a)
- |> Variable.declare_term (term_of D);
+ |> Variable.declare_term (Thm.term_of a)
+ |> Variable.declare_term (Thm.term_of D);
val vars = filter (fn ((v, _), T) => Type.could_unify (T, aT) andalso
not (member (op =) leave v)) (Term.add_vars (Thm.prop_of thm) []);
val c_vars = map (Proof_Context.cterm_of ctxt3 o Var) vars;
@@ -131,7 +131,7 @@
|> fold Simplifier.add_cong cong;
val thm' = thm
|> Drule.cterm_instantiate (c_vars ~~ c_exprs')
- |> fold_rev Thm.implies_intr (map cprop_of hyps)
+ |> fold_rev Thm.implies_intr (map Thm.cprop_of hyps)
|> Simplifier.asm_full_simplify simpset
in singleton (Variable.export ctxt5 ctxt1) thm' end;
--- a/src/HOL/Tools/lin_arith.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/lin_arith.ML Wed Mar 04 19:53:18 2015 +0100
@@ -62,8 +62,8 @@
fun mk_nat_thm thy t =
let
- val cn = cterm_of thy (Var (("n", 0), HOLogic.natT))
- and ct = cterm_of thy t
+ val cn = Thm.cterm_of thy (Var (("n", 0), HOLogic.natT))
+ and ct = Thm.cterm_of thy t
in Drule.instantiate_normalize ([], [(cn, ct)]) @{thm le0} end;
end;
@@ -288,7 +288,7 @@
(* checks if splitting with 'thm' is implemented *)
fun is_split_thm ctxt thm =
- (case concl_of thm of _ $ (_ $ (_ $ lhs) $ _) =>
+ (case Thm.concl_of thm of _ $ (_ $ (_ $ lhs) $ _) =>
(* Trueprop $ ((op =) $ (?P $ lhs) $ rhs) *)
(case head_of lhs of
Const (a, _) =>
--- a/src/HOL/Tools/nat_numeral_simprocs.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML Wed Mar 04 19:53:18 2015 +0100
@@ -216,10 +216,10 @@
val bal_add2 = @{thm nat_diff_add_eq2} RS trans
);
-fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (term_of ct)
-fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (term_of ct)
-fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (term_of ct)
-fun diff_cancel_numerals ctxt ct = DiffCancelNumerals.proc ctxt (term_of ct)
+fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (Thm.term_of ct)
+fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (Thm.term_of ct)
+fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (Thm.term_of ct)
+fun diff_cancel_numerals ctxt ct = DiffCancelNumerals.proc ctxt (Thm.term_of ct)
(*** Applying CombineNumeralsFun ***)
@@ -257,7 +257,7 @@
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
-fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (term_of ct)
+fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (Thm.term_of ct)
(*** Applying CancelNumeralFactorFun ***)
@@ -326,11 +326,11 @@
val neg_exchanges = true
)
-fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (term_of ct)
-fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (term_of ct)
-fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (term_of ct)
-fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (term_of ct)
-fun dvd_cancel_numeral_factor ctxt ct = DvdCancelNumeralFactor.proc ctxt (term_of ct)
+fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun dvd_cancel_numeral_factor ctxt ct = DvdCancelNumeralFactor.proc ctxt (Thm.term_of ct)
(*** Applying ExtractCommonTermFun ***)
@@ -405,10 +405,10 @@
fun simp_conv _ _ = SOME @{thm nat_mult_dvd_cancel_disj}
);
-fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (term_of ct)
-fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (term_of ct)
-fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (term_of ct)
-fun div_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (term_of ct)
-fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (term_of ct)
+fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (Thm.term_of ct)
+fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (Thm.term_of ct)
+fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (Thm.term_of ct)
+fun div_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (Thm.term_of ct)
+fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (Thm.term_of ct)
end;
--- a/src/HOL/Tools/numeral_simprocs.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML Wed Mar 04 19:53:18 2015 +0100
@@ -288,9 +288,9 @@
val bal_add2 = @{thm le_add_iff2} RS trans
);
-fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (term_of ct)
-fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (term_of ct)
-fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (term_of ct)
+fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (Thm.term_of ct)
+fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (Thm.term_of ct)
+fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (Thm.term_of ct)
structure CombineNumeralsData =
struct
@@ -350,9 +350,9 @@
structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData);
-fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (term_of ct)
+fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (Thm.term_of ct)
-fun field_combine_numerals ctxt ct = FieldCombineNumerals.proc ctxt (term_of ct)
+fun field_combine_numerals ctxt ct = FieldCombineNumerals.proc ctxt (Thm.term_of ct)
(** Constant folding for multiplication in semirings **)
@@ -368,7 +368,7 @@
structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data);
-fun assoc_fold ctxt ct = Semiring_Times_Assoc.proc ctxt (term_of ct)
+fun assoc_fold ctxt ct = Semiring_Times_Assoc.proc ctxt (Thm.term_of ct)
structure CancelNumeralFactorCommon =
struct
@@ -440,11 +440,11 @@
val neg_exchanges = true
)
-fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (term_of ct)
-fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (term_of ct)
-fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (term_of ct)
-fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (term_of ct)
-fun divide_cancel_numeral_factor ctxt ct = DivideCancelNumeralFactor.proc ctxt (term_of ct)
+fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (Thm.term_of ct)
+fun divide_cancel_numeral_factor ctxt ct = DivideCancelNumeralFactor.proc ctxt (Thm.term_of ct)
val field_divide_cancel_numeral_factor =
[prep_simproc @{theory}
@@ -577,19 +577,19 @@
fun simp_conv _ _ = SOME @{thm mult_divide_mult_cancel_left_if}
);
-fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (term_of ct)
-fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (term_of ct)
-fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (term_of ct)
-fun div_cancel_factor ctxt ct = DivCancelFactor.proc ctxt (term_of ct)
-fun mod_cancel_factor ctxt ct = ModCancelFactor.proc ctxt (term_of ct)
-fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (term_of ct)
-fun divide_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (term_of ct)
+fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (Thm.term_of ct)
+fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (Thm.term_of ct)
+fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (Thm.term_of ct)
+fun div_cancel_factor ctxt ct = DivCancelFactor.proc ctxt (Thm.term_of ct)
+fun mod_cancel_factor ctxt ct = ModCancelFactor.proc ctxt (Thm.term_of ct)
+fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (Thm.term_of ct)
+fun divide_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (Thm.term_of ct)
local
val zr = @{cpat "0"}
- val zT = ctyp_of_term zr
+ val zT = Thm.ctyp_of_term zr
val geq = @{cpat HOL.eq}
- val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd
+ val eqT = Thm.dest_ctyp (Thm.ctyp_of_term geq) |> hd
val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"}
val add_frac_num = mk_meta_eq @{thm "add_frac_num"}
val add_num_frac = mk_meta_eq @{thm "add_num_frac"}
@@ -608,8 +608,8 @@
let
val ((x,y),(w,z)) =
(Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct
- val _ = map (HOLogic.dest_number o term_of) [x,y,z,w]
- val T = ctyp_of_term x
+ val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z,w]
+ val T = Thm.ctyp_of_term x
val [y_nz, z_nz] = map (prove_nz ctxt T) [y, z]
val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq
in SOME (Thm.implies_elim (Thm.implies_elim th y_nz) z_nz)
@@ -619,17 +619,17 @@
fun proc2 phi ctxt ct =
let
val (l,r) = Thm.dest_binop ct
- val T = ctyp_of_term l
- in (case (term_of l, term_of r) of
+ val T = Thm.ctyp_of_term l
+ in (case (Thm.term_of l, Thm.term_of r) of
(Const(@{const_name Fields.divide},_)$_$_, _) =>
let val (x,y) = Thm.dest_binop l val z = r
- val _ = map (HOLogic.dest_number o term_of) [x,y,z]
+ val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z]
val ynz = prove_nz ctxt T y
in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz)
end
| (_, Const (@{const_name Fields.divide},_)$_$_) =>
let val (x,y) = Thm.dest_binop r val z = l
- val _ = map (HOLogic.dest_number o term_of) [x,y,z]
+ val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z]
val ynz = prove_nz ctxt T y
in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz)
end
@@ -640,50 +640,50 @@
fun is_number (Const(@{const_name Fields.divide},_)$a$b) = is_number a andalso is_number b
| is_number t = can HOLogic.dest_number t
- val is_number = is_number o term_of
+ val is_number = is_number o Thm.term_of
fun proc3 phi ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
Const(@{const_name Orderings.less},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
let
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"}
in SOME (mk_meta_eq th) end
| Const(@{const_name Orderings.less_eq},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
let
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"}
in SOME (mk_meta_eq th) end
| Const(@{const_name HOL.eq},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
let
val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"}
in SOME (mk_meta_eq th) end
| Const(@{const_name Orderings.less},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
let
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"}
in SOME (mk_meta_eq th) end
| Const(@{const_name Orderings.less_eq},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
let
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"}
in SOME (mk_meta_eq th) end
| Const(@{const_name HOL.eq},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
let
val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
val _ = map is_number [a,b,c]
- val T = ctyp_of_term c
+ val T = Thm.ctyp_of_term c
val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"}
in SOME (mk_meta_eq th) end
| _ => NONE)
--- a/src/HOL/Tools/record.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/record.ML Wed Mar 04 19:53:18 2015 +0100
@@ -77,8 +77,8 @@
let
fun match name ((name', _), _) = name = name';
fun getvar name =
- (case find_first (match name) (Term.add_vars (prop_of thm) []) of
- SOME var => cterm_of (theory_of_thm thm) (Var var)
+ (case find_first (match name) (Term.add_vars (Thm.prop_of thm) []) of
+ SOME var => Thm.cterm_of (Thm.theory_of_thm thm) (Var var)
| NONE => raise THM ("named_cterm_instantiate: " ^ name, 0, [thm]));
in
cterm_instantiate (map (apfst getvar) values) thm
@@ -97,7 +97,7 @@
let
val exists_thm =
UNIV_I
- |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global rep_type))] [];
+ |> Drule.instantiate' [SOME (Thm.ctyp_of thy (Logic.varifyT_global rep_type))] [];
val proj_constr = Abs_inverse OF [exists_thm];
val absT = Type (tyco, map TFree vs);
in
@@ -141,8 +141,8 @@
(*construct a type and body for the isomorphism constant by
instantiating the theorem to which the definition will be applied*)
val intro_inst =
- rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] iso_tuple_intro;
- val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst));
+ rep_inject RS named_cterm_instantiate [("abst", Thm.cterm_of typ_thy absC)] iso_tuple_intro;
+ val (_, body) = Logic.dest_equals (List.last (Thm.prems_of intro_inst));
val isomT = fastype_of body;
val isom_binding = Binding.suffix_name isoN b;
val isom_name = Sign.full_name typ_thy isom_binding;
@@ -1121,7 +1121,7 @@
fun get_upd_acc_cong_thm upd acc thy ss =
let
- val insts = [("upd", cterm_of thy upd), ("ac", cterm_of thy acc)];
+ val insts = [("upd", Thm.cterm_of thy upd), ("ac", Thm.cterm_of thy acc)];
val prop = Thm.concl_of (named_cterm_instantiate insts updacc_cong_triv);
in
Goal.prove (Proof_Context.init_global thy) [] [] prop
@@ -1364,7 +1364,7 @@
let
val thy = Proof_Context.theory_of ctxt;
- val goal = term_of cgoal;
+ val goal = Thm.term_of cgoal;
val frees = filter (is_recT o #2) (Term.add_frees goal []);
val has_rec = exists_Const
@@ -1375,9 +1375,9 @@
fun mk_split_free_tac free induct_thm i =
let
- val cfree = cterm_of thy free;
- val _$ (_ $ r) = concl_of induct_thm;
- val crec = cterm_of thy r;
+ val cfree = Thm.cterm_of thy free;
+ val _$ (_ $ r) = Thm.concl_of induct_thm;
+ val crec = Thm.cterm_of thy r;
val thm = cterm_instantiate [(crec, cfree)] induct_thm;
in
simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) i THEN
@@ -1415,7 +1415,7 @@
(*Split all records in the goal, which are quantified by !! or ALL.*)
fun split_tac ctxt = CSUBGOAL (fn (cgoal, i) =>
let
- val goal = term_of cgoal;
+ val goal = Thm.term_of cgoal;
val has_rec = exists_Const
(fn (s, Type (_, [Type (_, [T, _]), _])) =>
@@ -1466,12 +1466,12 @@
val params = Logic.strip_params g;
val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g);
val rule' = Thm.lift_rule cgoal rule;
- val (P, ys) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_assums_concl (prop_of rule')));
+ val (P, ys) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_assums_concl (Thm.prop_of rule')));
(*ca indicates if rule is a case analysis or induction rule*)
val (x, ca) =
(case rev (drop (length params) ys) of
[] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
- (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
+ (hd (rev (Logic.strip_assums_hyp (hd (Thm.prems_of rule')))))))), true)
| [x] => (head_of x, false));
val rule'' =
cterm_instantiate
@@ -1484,7 +1484,7 @@
| (_, T) :: _ =>
[(P, fold_rev Term.abs params (if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
(x, fold_rev Term.abs params (Bound 0))])) rule';
- in compose_tac ctxt (false, rule'', nprems_of rule) i end);
+ in compose_tac ctxt (false, rule'', Thm.nprems_of rule) i end);
fun extension_definition name fields alphas zeta moreT more vars thy =
@@ -1608,7 +1608,7 @@
(roughly) the definition of the accessor.*)
val surject = timeit_msg ctxt "record extension surjective proof:" (fn () =>
let
- val cterm_ext = cterm_of defs_thy ext;
+ val cterm_ext = Thm.cterm_of defs_thy ext;
val start = named_cterm_instantiate [("y", cterm_ext)] surject_assist_idE;
val tactic1 =
simp_tac (put_simpset HOL_basic_ss defs_ctxt addsimps [ext_def]) 1 THEN
@@ -1758,7 +1758,7 @@
fun mk_eq_refl thy =
@{thm equal_refl}
|> Thm.instantiate
- ([apply2 (ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], [])
+ ([apply2 (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], [])
|> Axclass.unoverload thy;
val ensure_random_record = ensure_sort_record (@{sort random}, mk_random_eq);
val ensure_exhaustive_record =
@@ -1945,8 +1945,8 @@
fun to_Var (Free (c, T)) = Var ((c, 1), T);
in mk_rec (map to_Var all_vars_more) 0 end;
- val cterm_rec = cterm_of ext_thy r_rec0;
- val cterm_vrs = cterm_of ext_thy r_rec0_Vars;
+ val cterm_rec = Thm.cterm_of ext_thy r_rec0;
+ val cterm_vrs = Thm.cterm_of ext_thy r_rec0_Vars;
val insts = [("v", cterm_rec), ("v'", cterm_vrs)];
val init_thm = named_cterm_instantiate insts updacc_eq_triv;
val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1;
@@ -2223,7 +2223,7 @@
sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs'
surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs';
- val sels = map (fst o Logic.dest_equals o prop_of) sel_defs;
+ val sels = map (fst o Logic.dest_equals o Thm.prop_of) sel_defs;
val final_thy =
thms_thy'
--- a/src/HOL/Tools/reification.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/reification.ML Wed Mar 04 19:53:18 2015 +0100
@@ -46,7 +46,7 @@
fun mk_congeq ctxt fs th =
let
- val Const (fN, _) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
+ val Const (fN, _) = th |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
|> fst |> strip_comb |> fst;
val cert = Proof_Context.cterm_of ctxt;
val ((_, [th']), ctxt') = Variable.import true [th] ctxt;
@@ -95,7 +95,7 @@
fun rearrange congs =
let
fun P (_, th) =
- let val @{term "Trueprop"} $ (Const (@{const_name HOL.eq}, _) $ l $ _) = concl_of th
+ let val @{term "Trueprop"} $ (Const (@{const_name HOL.eq}, _) $ l $ _) = Thm.concl_of th
in can dest_Var l end;
val (yes, no) = List.partition P congs;
in no @ yes end;
@@ -134,8 +134,8 @@
fun decomp_reify da cgns (ct, ctxt) bds =
let
val thy = Proof_Context.theory_of ctxt;
- val cert = cterm_of thy;
- val certT = ctyp_of thy;
+ val cert = Thm.cterm_of thy;
+ val certT = Thm.ctyp_of thy;
fun tryabsdecomp (ct, ctxt) bds =
(case Thm.term_of ct of
Abs (_, xT, ta) =>
@@ -167,7 +167,7 @@
(let
val (tyenv, tmenv) =
Pattern.match thy
- ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), Thm.term_of ct)
+ ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.concl_of cong), Thm.term_of ct)
(Vartab.empty, Vartab.empty);
val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
val (fts, its) =
@@ -189,15 +189,15 @@
fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation"
| tryeqs (eq :: eqs) (ct, ctxt) bds = ((
let
- val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
+ val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
val nths = get_nths rhs [];
val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
(insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
val thy = Proof_Context.theory_of ctxt'';
- val cert = cterm_of thy;
- val certT = ctyp_of thy;
+ val cert = Thm.cterm_of thy;
+ val certT = Thm.ctyp_of thy;
val vsns_map = vss ~~ vsns;
val xns_map = fst (split_list nths) ~~ xns;
val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
@@ -243,7 +243,7 @@
val tT = fastype_of (Thm.term_of ct);
fun isat eq =
let
- val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
+ val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
in exists_Const
(fn (n, ty) => n = @{const_name "List.nth"}
andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
@@ -256,7 +256,7 @@
fun mk_congs ctxt eqs =
let
- val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
+ val fs = fold_rev (fn eq => insert (op =) (eq |> Thm.prop_of |> HOLogic.dest_Trueprop
|> HOLogic.dest_eq |> fst |> strip_comb
|> fst)) eqs [];
val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
@@ -266,7 +266,7 @@
the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
fun prep_eq eq =
let
- val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
+ val (_, _ :: vs) = eq |> Thm.prop_of |> HOLogic.dest_Trueprop
|> HOLogic.dest_eq |> fst |> strip_comb;
val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
| _ => NONE) vs;
@@ -283,7 +283,7 @@
apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds);
fun is_list_var (Var (_, t)) = can dest_listT t
| is_list_var _ = false;
- val vars = th |> prop_of |> Logic.dest_equals |> snd
+ val vars = th |> Thm.prop_of |> Logic.dest_equals |> snd
|> strip_comb |> snd |> filter is_list_var;
val cert = Proof_Context.cterm_of ctxt;
val vs = map (fn v as Var (_, T) =>
--- a/src/HOL/Tools/sat.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/sat.ML Wed Mar 04 19:53:18 2015 +0100
@@ -73,7 +73,7 @@
val resolution_thm =
@{lemma "(P ==> False) ==> (~ P ==> False) ==> False" by (rule case_split)}
-val cP = cterm_of @{theory} (Var (("P", 0), HOLogic.boolT));
+val cP = Thm.cterm_of @{theory} (Var (("P", 0), HOLogic.boolT));
(* ------------------------------------------------------------------------- *)
(* lit_ord: an order on integers that considers their absolute values only, *)
@@ -354,7 +354,7 @@
(* [c_1, ..., c_n] |- c_1 && ... && c_n *)
val clauses_thm = Conjunction.intr_balanced (map Thm.assume sorted_clauses)
(* [c_1 && ... && c_n] |- c_1 && ... && c_n *)
- val cnf_cterm = cprop_of clauses_thm
+ val cnf_cterm = Thm.cprop_of clauses_thm
val cnf_thm = Thm.assume cnf_cterm
(* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *)
val cnf_clauses = Conjunction.elim_balanced (length sorted_clauses) cnf_thm
@@ -409,7 +409,7 @@
fun rawsat_tac ctxt i =
Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
- resolve_tac ctxt' [rawsat_thm ctxt' (map cprop_of prems)] 1) ctxt i;
+ resolve_tac ctxt' [rawsat_thm ctxt' (map Thm.cprop_of prems)] 1) ctxt i;
(* ------------------------------------------------------------------------- *)
(* pre_cnf_tac: converts the i-th subgoal *)
--- a/src/HOL/Tools/semiring_normalizer.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/semiring_normalizer.ML Wed Mar 04 19:53:18 2015 +0100
@@ -126,12 +126,12 @@
val field_funs =
let
fun numeral_is_const ct =
- case term_of ct of
+ case Thm.term_of ct of
Const (@{const_name Fields.divide},_) $ a $ b =>
can HOLogic.dest_number a andalso can HOLogic.dest_number b
| Const (@{const_name Fields.inverse},_)$t => can HOLogic.dest_number t
| t => can HOLogic.dest_number t
- fun dest_const ct = ((case term_of ct of
+ fun dest_const ct = ((case Thm.term_of ct of
Const (@{const_name Fields.divide},_) $ a $ b=>
Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
| Const (@{const_name Fields.inverse},_)$t =>
@@ -231,7 +231,7 @@
fun is_binop ct ct' =
(case Thm.term_of ct' of
- c $ _ $ _ => term_of ct aconv c
+ c $ _ $ _ => Thm.term_of ct aconv c
| _ => false);
fun dest_binop ct ct' =
@@ -240,7 +240,7 @@
fun inst_thm inst = Thm.instantiate ([], inst);
-val dest_number = term_of #> HOLogic.dest_number #> snd;
+val dest_number = Thm.term_of #> HOLogic.dest_number #> snd;
val is_number = can dest_number;
fun numeral01_conv ctxt =
@@ -838,7 +838,7 @@
addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps})
addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]);
-fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
+fun simple_cterm_ord t u = Term_Ord.term_ord (Thm.term_of t, Thm.term_of u) = LESS;
(* various normalizing conversions *)
--- a/src/HOL/Tools/set_comprehension_pointfree.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/set_comprehension_pointfree.ML Wed Mar 04 19:53:18 2015 +0100
@@ -430,7 +430,7 @@
in
HOLogic.Collect_const T $ absdummy T (list_ex vs (HOLogic.mk_conj (eq, t'')))
end;
- fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (prop_of th))
+ fun is_eq th = is_some (try (HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.prop_of th))
val unfold_thms = @{thms split_paired_all mem_Collect_eq prod.case}
fun tac ctxt =
resolve_tac ctxt @{thms set_eqI}
@@ -446,10 +446,10 @@
simp_tac (put_simpset HOL_basic_ss ctxt addsimps (filter is_eq prems)) 1) ctxt
THEN' TRY o assume_tac ctxt
in
- case try mk_term (term_of ct) of
+ case try mk_term (Thm.term_of ct) of
NONE => Thm.reflexive ct
| SOME t' =>
- Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (term_of ct, t')))
+ Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (Thm.term_of ct, t')))
(fn {context, ...} => tac context 1)
RS @{thm eq_reflection}
end
@@ -473,7 +473,7 @@
Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))
(empty_simpset ctxt' addsimps thms)
val prep_eq = (comprehension_conv ctxt' then_conv unfold_conv prep_thms) ct
- val t'' = term_of (Thm.rhs_of prep_eq)
+ val t'' = Thm.term_of (Thm.rhs_of prep_eq)
fun mk_thm (fm, t''') = Goal.prove ctxt' [] []
(HOLogic.mk_Trueprop (HOLogic.mk_eq (t'', t'''))) (fn {context, ...} => tac context fm 1)
fun unfold th = th RS ((prep_eq RS meta_eq_to_obj_eq) RS @{thm trans})
@@ -487,7 +487,7 @@
fun base_simproc ctxt redex =
let
- val set_compr = term_of redex
+ val set_compr = Thm.term_of redex
in
conv ctxt set_compr
|> Option.map (fn thm => thm RS @{thm eq_reflection})
@@ -496,14 +496,14 @@
fun instantiate_arg_cong ctxt pred =
let
val arg_cong = Thm.incr_indexes (maxidx_of_term pred + 1) @{thm arg_cong}
- val f $ _ = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of arg_cong)))
+ val f $ _ = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.concl_of arg_cong)))
in
cterm_instantiate [(Proof_Context.cterm_of ctxt f, Proof_Context.cterm_of ctxt pred)] arg_cong
end;
fun simproc ctxt redex =
let
- val pred $ set_compr = term_of redex
+ val pred $ set_compr = Thm.term_of redex
val arg_cong' = instantiate_arg_cong ctxt pred
in
conv ctxt set_compr
--- a/src/HOL/Tools/simpdata.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/simpdata.ML Wed Mar 04 19:53:18 2015 +0100
@@ -41,12 +41,13 @@
fun mk_meta_eq r = r RS @{thm eq_reflection};
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;
-fun mk_eq th = case concl_of th
+fun mk_eq th =
+ (case Thm.concl_of th of
(*expects Trueprop if not == *)
- of Const (@{const_name Pure.eq},_) $ _ $ _ => th
- | _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => mk_meta_eq th
- | _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI}
- | _ => th RS @{thm Eq_TrueI}
+ Const (@{const_name Pure.eq},_) $ _ $ _ => th
+ | _ $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => mk_meta_eq th
+ | _ $ (Const (@{const_name Not}, _) $ _) => th RS @{thm Eq_FalseI}
+ | _ => th RS @{thm Eq_TrueI})
fun mk_eq_True (_: Proof.context) r =
SOME (r RS @{thm meta_eq_to_obj_eq} RS @{thm Eq_TrueI}) handle Thm.THM _ => NONE;
@@ -87,7 +88,7 @@
resolve_tac ctxt [lift_meta_eq_to_obj_eq i st] i st) rl)
in
mk_meta_eq rl' handle THM _ =>
- if can Logic.dest_equals (concl_of rl') then rl'
+ if can Logic.dest_equals (Thm.concl_of rl') then rl'
else error "Conclusion of congruence rules must be =-equality"
end |> zero_var_indexes;
@@ -101,7 +102,7 @@
if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls
else Variable.trade (K (fn [thm'] => res thm' rls)) thm_ctxt [thm];
in
- case concl_of thm
+ case Thm.concl_of thm
of Const (@{const_name Trueprop}, _) $ p => (case head_of p
of Const (a, _) => (case AList.lookup (op =) pairs a
of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm])
--- a/src/HOL/Tools/split_rule.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/split_rule.ML Wed Mar 04 19:53:18 2015 +0100
@@ -89,7 +89,7 @@
(*curries ALL function variables occurring in a rule's conclusion*)
fun split_rule ctxt rl =
- fold_rev split_rule_var' (Misc_Legacy.term_vars (concl_of rl)) rl
+ fold_rev split_rule_var' (Misc_Legacy.term_vars (Thm.concl_of rl)) rl
|> remove_internal_split ctxt
|> Drule.export_without_context;
--- a/src/HOL/Tools/try0.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Tools/try0.ML Wed Mar 04 19:53:18 2015 +0100
@@ -29,7 +29,7 @@
(case (case timeout_opt of
SOME timeout => TimeLimit.timeLimit timeout
| NONE => fn f => fn x => f x) (Seq.pull o tac) (pre st) of
- SOME (x, _) => nprems_of (post x) < nprems_of goal
+ SOME (x, _) => Thm.nprems_of (post x) < Thm.nprems_of goal
| NONE => false)
end;
@@ -68,7 +68,7 @@
let val attrs = attrs_text attrs quad in
apply_generic timeout_opt name
((name ^ attrs |> attrs <> "" ? enclose "(" ")") ^
- (if all_goals andalso nprems_of (#goal (Proof.goal st)) > 1 then "[1]" else ""))
+ (if all_goals andalso Thm.nprems_of (#goal (Proof.goal st)) > 1 then "[1]" else ""))
I (#goal o Proof.goal)
(apply_named_method_on_first_goal (Proof.context_of st) (name ^ attrs)) st
end
@@ -141,7 +141,7 @@
| Try => "Try0 found a proof"
| Normal => "Try this") ^ ": " ^
Active.sendback_markup [Markup.padding_command]
- ((if nprems_of (#goal (Proof.goal st)) = 1 then "by"
+ ((if Thm.nprems_of (#goal (Proof.goal st)) = 1 then "by"
else "apply") ^ " " ^ command) ^
(case xs of
[(_, ms)] => " (" ^ time_string ms ^ ")."
--- a/src/HOL/Topological_Spaces.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Topological_Spaces.thy Wed Mar 04 19:53:18 2015 +0100
@@ -451,7 +451,7 @@
(@{thm allI} RS @{thm always_eventually})
|> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms
|> fold (fn _ => fn thm => @{thm impI} RS thm) thms
- val cases_prop = prop_of (raw_elim_thm RS st)
+ val cases_prop = Thm.prop_of (raw_elim_thm RS st)
val cases = (Rule_Cases.make_common (thy, cases_prop) [(("elim", []), [])])
in
CASES cases (rtac raw_elim_thm 1)
--- a/src/HOL/Word/WordBitwise.thy Tue Mar 03 19:08:04 2015 +0100
+++ b/src/HOL/Word/WordBitwise.thy Wed Mar 04 19:53:18 2015 +0100
@@ -507,7 +507,7 @@
@{cterm "[] :: nat list"} ns;
fun upt_conv ctxt ct =
- case term_of ct of
+ case Thm.term_of ct of
(@{const upt} $ n $ m) =>
let
val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m);
@@ -529,7 +529,7 @@
proc = K upt_conv};
fun word_len_simproc_fn ctxt ct =
- case term_of ct of
+ case Thm.term_of ct of
Const (@{const_name len_of}, _) $ t => (let
val T = fastype_of t |> dest_Type |> snd |> the_single
val n = Numeral.mk_cnumber @{ctyp nat} (Word_Lib.dest_binT T);
@@ -551,7 +551,7 @@
fun nat_get_Suc_simproc_fn n_sucs ctxt ct =
let
val thy = Proof_Context.theory_of ctxt;
- val (f $ arg) = (term_of ct);
+ val (f $ arg) = Thm.term_of ct;
val n = (case arg of @{term nat} $ n => n | n => n)
|> HOLogic.dest_number |> snd;
val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs)
@@ -560,7 +560,7 @@
(replicate i @{term Suc});
val _ = if arg = arg' then raise TERM ("", []) else ();
fun propfn g = HOLogic.mk_eq (g arg, g arg')
- |> HOLogic.mk_Trueprop |> cterm_of thy;
+ |> HOLogic.mk_Trueprop |> Thm.cterm_of thy;
val eq1 = Goal.prove_internal ctxt [] (propfn I)
(K (simp_tac (put_simpset word_ss ctxt) 1));
in Goal.prove_internal ctxt [] (propfn (curry (op $) f))
--- a/src/Provers/Arith/cancel_div_mod.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/Arith/cancel_div_mod.ML Wed Mar 04 19:53:18 2015 +0100
@@ -72,7 +72,7 @@
fun proc ctxt ct =
let
- val t = term_of ct;
+ val t = Thm.term_of ct;
val (divs, mods) = coll_div_mod t ([], []);
in
if null divs orelse null mods then NONE
--- a/src/Provers/Arith/fast_lin_arith.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/Arith/fast_lin_arith.ML Wed Mar 04 19:53:18 2015 +0100
@@ -882,7 +882,7 @@
| elim_neq (neqs as (neq :: _)) (asms', asm::asms) =
(case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) [neq] of
SOME spl =>
- let val (ct1, ct2) = extract (cprop_of spl)
+ let val (ct1, ct2) = extract (Thm.cprop_of spl)
val thm1 = Thm.assume ct1
val thm2 = Thm.assume ct2
in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]),
@@ -905,7 +905,7 @@
let
val thy = Proof_Context.theory_of ctxt
val nTconcl = LA_Logic.neg_prop Tconcl
- val cnTconcl = cterm_of thy nTconcl
+ val cnTconcl = Thm.cterm_of thy nTconcl
val nTconclthm = Thm.assume cnTconcl
val tree = (if split_neq then splitasms ctxt else Tip) (thms @ [nTconclthm])
val (Falsethm, _) = fwdproof ctxt tree js
--- a/src/Provers/clasimp.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/clasimp.ML Wed Mar 04 19:53:18 2015 +0100
@@ -65,7 +65,7 @@
fun add_iff safe unsafe =
Thm.declaration_attribute (fn th =>
let
- val n = nprems_of th;
+ val n = Thm.nprems_of th;
val (elim, intro) = if n = 0 then safe else unsafe;
val zero_rotate = zero_var_indexes o rotate_prems n;
in
@@ -77,7 +77,7 @@
end);
fun del_iff del = Thm.declaration_attribute (fn th =>
- let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in
+ let val zero_rotate = zero_var_indexes o rotate_prems (Thm.nprems_of th) in
Thm.attribute_declaration del (zero_rotate (th RS Data.iffD2)) #>
Thm.attribute_declaration del (Tactic.make_elim (zero_rotate (th RS Data.iffD1)))
handle THM _ =>
--- a/src/Provers/classical.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/classical.ML Wed Mar 04 19:53:18 2015 +0100
@@ -351,7 +351,7 @@
let
val th' = classical_rule (flat_rule context th);
val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)
- List.partition (fn rl => nprems_of rl=1) [th'];
+ List.partition (fn rl => Thm.nprems_of rl=1) [th'];
val nI = Item_Net.length safeIs;
val nE = Item_Net.length safeEs + 1;
val _ = warn_claset context th cs;
@@ -467,7 +467,7 @@
if Item_Net.member safeEs th then
let
val th' = classical_rule (flat_rule context th);
- val (safe0_rls, safep_rls) = List.partition (fn rl => nprems_of rl = 1) [th'];
+ val (safe0_rls, safep_rls) = List.partition (fn rl => Thm.nprems_of rl = 1) [th'];
in
CS
{safe0_netpair = delete ([], safe0_rls) safe0_netpair,
--- a/src/Provers/hypsubst.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/hypsubst.ML Wed Mar 04 19:53:18 2015 +0100
@@ -184,14 +184,14 @@
(case (if b then t else t') of
Bound j => subst_bounds (map Bound ((1 upto j) @ 0 :: (j + 2 upto length ps)), Q)
| t => Term.abstract_over (t, Term.incr_boundvars 1 Q));
- val (instT, _) = Thm.match (apply2 (cterm_of thy o Logic.mk_type) (V, U));
+ val (instT, _) = Thm.match (apply2 (Thm.cterm_of thy o Logic.mk_type) (V, U));
in
compose_tac ctxt (true, Drule.instantiate_normalize (instT,
- map (apply2 (cterm_of thy))
+ map (apply2 (Thm.cterm_of thy))
[(Var (ixn, Ts ---> U --> body_type T), u),
(Var (fst (dest_Var (head_of v1)), Ts ---> U), fold_rev Term.abs ps t),
(Var (fst (dest_Var (head_of v2)), Ts ---> U), fold_rev Term.abs ps t')]) rl',
- nprems_of rl) i
+ Thm.nprems_of rl) i
end
| NONE => no_tac);
@@ -256,7 +256,7 @@
| imptac (r, hyp::hyps) st =
let
val (hyp', _) =
- term_of (Thm.cprem_of st i)
+ Thm.term_of (Thm.cprem_of st i)
|> Logic.strip_assums_concl
|> Data.dest_Trueprop |> Data.dest_imp;
val (r', tac) =
--- a/src/Provers/order.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/order.ML Wed Mar 04 19:53:18 2015 +0100
@@ -1235,7 +1235,9 @@
let
val thy = Proof_Context.theory_of ctxt;
val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
- val Hs = map prop_of prems @ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
+ val Hs =
+ map Thm.prop_of prems @
+ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
val lesss = flat (map_index (mkasm decomp less_thms thy) Hs)
val prfs = gen_solve mkconcl thy (lesss, C);
--- a/src/Provers/quantifier1.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/quantifier1.ML Wed Mar 04 19:53:18 2015 +0100
@@ -181,7 +181,7 @@
in quant xs (qC $ Abs (x, T, Q)) end;
fun rearrange_all ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
F as (all as Const (q, _)) $ Abs (x, T, P) =>
(case extract_quant extract_imp q P of
NONE => NONE
@@ -191,7 +191,7 @@
| _ => NONE);
fun rearrange_ball tac ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
F as Ball $ A $ Abs (x, T, P) =>
(case extract_imp true [] P of
NONE => NONE
@@ -203,7 +203,7 @@
| _ => NONE);
fun rearrange_ex ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
F as (ex as Const (q, _)) $ Abs (x, T, P) =>
(case extract_quant extract_conj q P of
NONE => NONE
@@ -213,7 +213,7 @@
| _ => NONE);
fun rearrange_bex tac ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
F as Bex $ A $ Abs (x, T, P) =>
(case extract_conj true [] P of
NONE => NONE
@@ -223,7 +223,7 @@
| _ => NONE);
fun rearrange_Collect tac ctxt ct =
- (case term_of ct of
+ (case Thm.term_of ct of
F as Collect $ Abs (x, T, P) =>
(case extract_conj true [] P of
NONE => NONE
--- a/src/Provers/splitter.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/splitter.ML Wed Mar 04 19:53:18 2015 +0100
@@ -55,17 +55,19 @@
fun split_format_err () = error "Wrong format for split rule";
-fun split_thm_info thm = case concl_of (Data.mk_eq thm) of
- Const(@{const_name Pure.eq}, _) $ (Var _ $ t) $ c => (case strip_comb t of
- (Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false)
- | _ => split_format_err ())
- | _ => split_format_err ();
+fun split_thm_info thm =
+ (case Thm.concl_of (Data.mk_eq thm) of
+ Const(@{const_name Pure.eq}, _) $ (Var _ $ t) $ c =>
+ (case strip_comb t of
+ (Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false)
+ | _ => split_format_err ())
+ | _ => split_format_err ());
fun cmap_of_split_thms thms =
let
val splits = map Data.mk_eq thms
fun add_thm thm cmap =
- (case concl_of thm of _ $ (t as _ $ lhs) $ _ =>
+ (case Thm.concl_of thm of _ $ (t as _ $ lhs) $ _ =>
(case strip_comb lhs of (Const(a,aT),args) =>
let val info = (aT,lhs,thm,fastype_of t,length args)
in case AList.lookup (op =) cmap a of
@@ -102,7 +104,7 @@
(fn {context = ctxt, prems} =>
rewrite_goals_tac ctxt prems THEN resolve_tac ctxt [reflexive_thm] 1)
-val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = prop_of lift;
+val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = Thm.prop_of lift;
val trlift = lift RS transitive_thm;
@@ -286,7 +288,7 @@
fun select cmap state i =
let
val thy = Thm.theory_of_thm state
- val goal = term_of (Thm.cprem_of state i);
+ val goal = Thm.term_of (Thm.cprem_of state i);
val Ts = rev (map #2 (Logic.strip_params goal));
val _ $ t $ _ = Logic.strip_assums_concl goal;
in (Ts, t, sort shorter (split_posns cmap thy Ts t)) end;
@@ -313,7 +315,7 @@
fun inst_lift Ts t (T, U, pos) state i =
let
- val cert = cterm_of (Thm.theory_of_thm state);
+ val cert = Thm.cterm_of (Thm.theory_of_thm state);
val (cntxt, u) = mk_cntxt t pos (T --> U);
val trlift' = Thm.lift_rule (Thm.cprem_of state i)
(Thm.rename_boundvars abs_lift u trlift);
@@ -341,7 +343,7 @@
val thm' = Thm.lift_rule (Thm.cprem_of state i) thm;
val (P, _) = strip_comb (fst (Logic.dest_equals
(Logic.strip_assums_concl (Thm.prop_of thm'))));
- val cert = cterm_of (Thm.theory_of_thm state);
+ val cert = Thm.cterm_of (Thm.theory_of_thm state);
val cntxt = mk_cntxt_splitthm t tt TB;
in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm'
end;
--- a/src/Provers/trancl.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Provers/trancl.ML Wed Mar 04 19:53:18 2015 +0100
@@ -93,8 +93,8 @@
fun prove thy r asms =
let
fun inst thm =
- let val SOME (_, _, r', _) = decomp (concl_of thm)
- in Drule.cterm_instantiate [(cterm_of thy r', cterm_of thy r)] thm end;
+ let val SOME (_, _, r', _) = decomp (Thm.concl_of thm)
+ in Drule.cterm_instantiate [(Thm.cterm_of thy r', Thm.cterm_of thy r)] thm end;
fun pr (Asm i) = nth asms i
| pr (Thm (prfs, thm)) = map pr prfs MRS inst thm;
in pr end;
@@ -543,7 +543,7 @@
in
Subgoal.FOCUS (fn {context = ctxt', prems, concl, ...} =>
let
- val SOME (_, _, rel', _) = decomp (term_of concl);
+ val SOME (_, _, rel', _) = decomp (Thm.term_of concl);
val thms = map (prove thy rel' prems) prfs
in resolve_tac ctxt' [prove thy rel' thms prf] 1 end) ctxt n st
end
@@ -562,7 +562,7 @@
in
Subgoal.FOCUS (fn {context = ctxt', prems, concl, ...} =>
let
- val SOME (_, _, rel', _) = decomp (term_of concl);
+ val SOME (_, _, rel', _) = decomp (Thm.term_of concl);
val thms = map (prove thy rel' prems) prfs
in resolve_tac ctxt' [prove thy rel' thms prf] 1 end) ctxt n st
end
--- a/src/Pure/Isar/expression.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Isar/expression.ML Wed Mar 04 19:53:18 2015 +0100
@@ -706,7 +706,7 @@
fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy =
let
val ctxt = Proof_Context.init_global thy;
- val defs' = map (cterm_of thy #> Assumption.assume ctxt #> Drule.abs_def) defs;
+ val defs' = map (Thm.cterm_of thy #> Assumption.assume ctxt #> Drule.abs_def) defs;
val (a_pred, a_intro, a_axioms, thy'') =
if null exts then (NONE, NONE, [], thy)
@@ -807,7 +807,7 @@
val notes =
if is_some asm then
[("", [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []),
- [([Assumption.assume pred_ctxt (cterm_of thy' (the asm))],
+ [([Assumption.assume pred_ctxt (Thm.cterm_of thy' (the asm))],
[(Attrib.internal o K) Locale.witness_add])])])]
else [];
--- a/src/Pure/Isar/proof.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Isar/proof.ML Wed Mar 04 19:53:18 2015 +0100
@@ -474,7 +474,8 @@
val thy = Proof_Context.theory_of ctxt;
val _ =
- Theory.subthy (theory_of_thm goal, thy) orelse error "Bad background theory of goal state";
+ Theory.subthy (Thm.theory_of_thm goal, thy) orelse
+ error "Bad background theory of goal state";
val _ = Thm.no_prems goal orelse error (Proof_Display.string_of_goal ctxt goal);
fun lost_structure () = error ("Lost goal structure:\n" ^ Display.string_of_thm ctxt goal);
--- a/src/Pure/Isar/rule_cases.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Isar/rule_cases.ML Wed Mar 04 19:53:18 2015 +0100
@@ -195,7 +195,7 @@
fun (tac1 THEN_ALL_NEW_CASES tac2) i st =
st |> tac1 i |> Seq.maps (fn (cases, st') =>
- CASES cases (Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st)) st');
+ CASES cases (Seq.INTERVAL tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st)) st');
--- a/src/Pure/Isar/runtime.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Isar/runtime.ML Wed Mar 04 19:53:18 2015 +0100
@@ -112,7 +112,7 @@
raised exn "TERM" (msg :: robust_context context Syntax.string_of_term ts)
| CTERM (msg, cts) =>
raised exn "CTERM"
- (msg :: robust_context context Syntax.string_of_term (map term_of cts))
+ (msg :: robust_context context Syntax.string_of_term (map Thm.term_of cts))
| THM (msg, i, thms) =>
raised exn ("THM " ^ string_of_int i)
(msg :: robust_context context Display.string_of_thm thms)
--- a/src/Pure/Proof/extraction.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Proof/extraction.ML Wed Mar 04 19:53:18 2015 +0100
@@ -319,8 +319,8 @@
let
val name = Thm.derivation_name thm;
val _ = name <> "" orelse error "add_realizers: unnamed theorem";
- val prop = Thm.unconstrainT thm |> prop_of |>
- Pattern.rewrite_term thy' (map (Logic.dest_equals o prop_of) defs) [];
+ val prop = Thm.unconstrainT thm |> Thm.prop_of |>
+ Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) [];
val vars = vars_of prop;
val vars' = filter_out (fn v =>
member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars;
@@ -354,7 +354,7 @@
val procs = maps (rev o fst o snd) types;
val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns);
val prop' = Pattern.rewrite_term thy'
- (map (Logic.dest_equals o prop_of) defs) [] prop;
+ (map (Logic.dest_equals o Thm.prop_of) defs) [] prop;
in freeze_thaw (condrew thy' eqns
(procs @ [typeof_proc [] vs, rlz_proc]))
(Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop')
@@ -364,8 +364,8 @@
let
val S = Sign.defaultS thy;
val ((atyp_map, constraints, _), prop') =
- Logic.unconstrainT (#shyps (rep_thm thm)) (prop_of thm);
- val atyps = fold_types (fold_atyps (insert (op =))) (prop_of thm) [];
+ Logic.unconstrainT (#shyps (Thm.rep_thm thm)) (Thm.prop_of thm);
+ val atyps = fold_types (fold_atyps (insert (op =))) (Thm.prop_of thm) [];
val Ts = map_filter (fn ((v, i), _) => if member (op =) vs v then
SOME (TVar (("'" ^ v, i), [])) else NONE)
(rev (Term.add_vars prop' []));
@@ -397,7 +397,7 @@
(if is_def then
{realizes_eqns = realizes_eqns,
typeof_eqns = add_rule ([], Logic.dest_equals (map_types
- Type.strip_sorts (prop_of (Drule.abs_def thm)))) typeof_eqns,
+ Type.strip_sorts (Thm.prop_of (Drule.abs_def thm)))) typeof_eqns,
types = types,
realizers = realizers, defs = insert Thm.eq_thm thm defs,
expand = expand, prep = prep}
--- a/src/Pure/Proof/proof_checker.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Proof/proof_checker.ML Wed Mar 04 19:53:18 2015 +0100
@@ -38,7 +38,7 @@
fun pretty_prf thy vs Hs prf =
let val prf' = prf |> Proofterm.prf_subst_bounds (map Free vs) |>
- Proofterm.prf_subst_pbounds (map (Hyp o prop_of) Hs)
+ Proofterm.prf_subst_pbounds (map (Hyp o Thm.prop_of) Hs)
in
(Proof_Syntax.pretty_proof (Syntax.init_pretty_global thy) prf',
Syntax.pretty_term_global thy (Reconstruct.prop_of prf'))
--- a/src/Pure/Proof/proof_rewrite_rules.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Proof/proof_rewrite_rules.ML Wed Mar 04 19:53:18 2015 +0100
@@ -253,7 +253,7 @@
fun elim_defs thy r defs prf =
let
val defs' = map (Logic.dest_equals o
- map_types Type.strip_sorts o prop_of o Drule.abs_def) defs;
+ map_types Type.strip_sorts o Thm.prop_of o Drule.abs_def) defs;
val defnames = map Thm.derivation_name defs;
val f = if not r then I else
let
--- a/src/Pure/Proof/proof_syntax.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Proof/proof_syntax.ML Wed Mar 04 19:53:18 2015 +0100
@@ -197,7 +197,7 @@
|> add_proof_atom_consts
(map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names);
in
- (cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of)
+ (Thm.cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of)
end;
fun read_term thy topsort =
--- a/src/Pure/Tools/rule_insts.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/Tools/rule_insts.ML Wed Mar 04 19:53:18 2015 +0100
@@ -213,7 +213,7 @@
(* Tactic *)
fun tac i st = CSUBGOAL (fn (cgoal, _) =>
let
- val goal = term_of cgoal;
+ val goal = Thm.term_of cgoal;
val params = Logic.strip_params goal; (*params of subgoal i as string typ pairs*)
val params = rev (Term.rename_wrt_term goal params)
(*as they are printed: bound variables with*)
@@ -265,12 +265,12 @@
fun liftterm t =
fold_rev absfree (param_names ~~ paramTs) (Logic.incr_indexes (paramTs, inc) t);
fun liftpair (cv, ct) = (cterm_fun liftvar cv, cterm_fun liftterm ct);
- val lifttvar = apply2 (ctyp_of thy o Logic.incr_tvar inc);
+ val lifttvar = apply2 (Thm.ctyp_of thy o Logic.incr_tvar inc);
val rule = Drule.instantiate_normalize
(map lifttvar envT', map liftpair cenv)
(Thm.lift_rule cgoal thm)
in
- compose_tac ctxt' (bires_flag, rule, nprems_of thm) i
+ compose_tac ctxt' (bires_flag, rule, Thm.nprems_of thm) i
end) i st;
in tac end;
--- a/src/Pure/axclass.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/axclass.ML Wed Mar 04 19:53:18 2015 +0100
@@ -191,7 +191,7 @@
thy2
|> update_classrel ((c1, c2),
(the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2))
- |> Drule.instantiate' [SOME (ctyp_of thy2 (TVar ((Name.aT, 0), [])))] []
+ |> Drule.instantiate' [SOME (Thm.ctyp_of thy2 (TVar ((Name.aT, 0), [])))] []
|> Thm.close_derivation));
val proven = is_classrel thy1;
@@ -228,7 +228,7 @@
c1 = c2 andalso Sorts.sorts_le algebra (Ss2, Ss)) ars);
val names = Name.invent Name.context Name.aT (length Ss);
- val std_vars = map (fn a => SOME (ctyp_of thy (TVar ((a, 0), [])))) names;
+ val std_vars = map (fn a => SOME (Thm.ctyp_of thy (TVar ((a, 0), [])))) names;
val completions = super_class_completions |> map (fn c1 =>
let
@@ -396,7 +396,7 @@
val (th', thy') = Global_Theory.store_thm (binding, th) thy;
val th'' = th'
|> Thm.unconstrainT
- |> Drule.instantiate' [SOME (ctyp_of thy' (TVar ((Name.aT, 0), [])))] [];
+ |> Drule.instantiate' [SOME (Thm.ctyp_of thy' (TVar ((Name.aT, 0), [])))] [];
in
thy'
|> Sign.primitive_classrel (c1, c2)
@@ -418,7 +418,7 @@
val args = Name.invent_names Name.context Name.aT Ss;
val T = Type (t, map TFree args);
- val std_vars = map (fn (a, S) => SOME (ctyp_of thy' (TVar ((a, 0), [])))) args;
+ val std_vars = map (fn (a, S) => SOME (Thm.ctyp_of thy' (TVar ((a, 0), [])))) args;
val missing_params = Sign.complete_sort thy' [c]
|> maps (these o Option.map #params o try (get_info thy'))
--- a/src/Pure/drule.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/drule.ML Wed Mar 04 19:53:18 2015 +0100
@@ -130,7 +130,7 @@
| _ => ct);
(*The premises of a theorem, as a cterm list*)
-val cprems_of = strip_imp_prems o cprop_of;
+val cprems_of = strip_imp_prems o Thm.cprop_of;
fun cterm_fun f ct = Thm.cterm_of (Thm.theory_of_cterm ct) (f (Thm.term_of ct));
fun ctyp_fun f cT = Thm.ctyp_of (Thm.theory_of_ctyp cT) (f (Thm.typ_of cT));
@@ -168,7 +168,7 @@
(*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs
of the meta-equality returned by the beta_conversion rule.*)
fun beta_conv x y =
- Thm.dest_arg (cprop_of (Thm.beta_conversion false (Thm.apply x y)));
+ Thm.dest_arg (Thm.cprop_of (Thm.beta_conversion false (Thm.apply x y)));
@@ -406,8 +406,8 @@
fun extensional eq =
let val eq' =
- Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (cprop_of eq)))) eq
- in Thm.equal_elim (Thm.eta_conversion (cprop_of eq')) eq' end;
+ Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (Thm.cprop_of eq)))) eq
+ in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end;
val equals_cong =
store_standard_thm_open (Binding.make ("equals_cong", @{here}))
@@ -452,7 +452,7 @@
fun binop_cong_rule ct th1 th2 = Thm.combination (arg_cong_rule ct th1) th2;
local
- val dest_eq = Thm.dest_equals o cprop_of
+ val dest_eq = Thm.dest_equals o Thm.cprop_of
val rhs_of = snd o dest_eq
in
fun beta_eta_conversion t =
@@ -467,7 +467,7 @@
(*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*)
fun eta_contraction_rule th =
- Thm.equal_elim (Thm.eta_conversion (cprop_of th)) th;
+ Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of th)) th;
(* abs_def *)
@@ -482,7 +482,7 @@
fun contract_lhs th =
Thm.transitive (Thm.symmetric (beta_eta_conversion
- (fst (Thm.dest_equals (cprop_of th))))) th;
+ (fst (Thm.dest_equals (Thm.cprop_of th))))) th;
fun var_args ct =
(case try Thm.dest_comb ct of
@@ -793,7 +793,7 @@
| cterm_instantiate ctpairs th =
let
val (thy, tye, _) = fold_rev add_types ctpairs (Thm.theory_of_thm th, Vartab.empty, 0);
- val certT = ctyp_of thy;
+ val certT = Thm.ctyp_of thy;
val instT =
Vartab.fold (fn (xi, (S, T)) =>
cons (certT (TVar (xi, S)), certT (Envir.norm_type tye T))) tye [];
--- a/src/Pure/goal.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/goal.ML Wed Mar 04 19:53:18 2015 +0100
@@ -214,7 +214,8 @@
| SOME st =>
let
val _ =
- Theory.subthy (theory_of_thm st, thy) orelse err "Bad background theory of goal state";
+ Theory.subthy (Thm.theory_of_thm st, thy) orelse
+ err "Bad background theory of goal state";
val res =
(finish ctxt' st
|> Drule.flexflex_unique (SOME ctxt')
--- a/src/Pure/more_thm.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/more_thm.ML Wed Mar 04 19:53:18 2015 +0100
@@ -273,8 +273,8 @@
);
fun declare_hyps ct ctxt =
- if Theory.subthy (theory_of_cterm ct, Proof_Context.theory_of ctxt) then
- (Hyps.map o apfst) (Termtab.update (term_of ct, ())) ctxt
+ if Theory.subthy (Thm.theory_of_cterm ct, Proof_Context.theory_of ctxt) then
+ (Hyps.map o apfst) (Termtab.update (Thm.term_of ct, ())) ctxt
else raise CTERM ("assume_hyps: bad background theory", [ct]);
fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt);
--- a/src/Pure/raw_simplifier.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/raw_simplifier.ML Wed Mar 04 19:53:18 2015 +0100
@@ -182,11 +182,11 @@
end;
fun rrule_extra_vars elhs thm =
- rewrite_rule_extra_vars [] (term_of elhs) (Thm.full_prop_of thm);
+ rewrite_rule_extra_vars [] (Thm.term_of elhs) (Thm.full_prop_of thm);
fun mk_rrule2 {thm, name, lhs, elhs, perm} =
let
- val t = term_of elhs;
+ val t = Thm.term_of elhs;
val fo = Pattern.first_order t orelse not (Pattern.pattern t);
val extra = rrule_extra_vars elhs thm;
in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end;
@@ -461,7 +461,7 @@
fun del_rrule (rrule as {thm, elhs, ...}) ctxt =
ctxt |> map_simpset1 (fn (rules, prems, depth) =>
- (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, depth))
+ (Net.delete_term eq_rrule (Thm.term_of elhs, rrule) rules, prems, depth))
handle Net.DELETE =>
(cond_warning ctxt (fn () => print_thm ctxt "Rewrite rule not in simpset:" ("", thm)); ctxt);
@@ -470,7 +470,7 @@
ctxt |> map_simpset1 (fn (rules, prems, depth) =>
let
val rrule2 as {elhs, ...} = mk_rrule2 rrule;
- val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules;
+ val rules' = Net.insert_term eq_rrule (Thm.term_of elhs, rrule2) rules;
in (rules', prems, depth) end)
handle Net.INSERT =>
(cond_warning ctxt (fn () => print_thm ctxt "Ignoring duplicate rewrite rule:" ("", thm));
@@ -496,12 +496,12 @@
val (lhs, rhs) = Thm.dest_equals concl handle TERM _ =>
raise SIMPLIFIER ("Rewrite rule not a meta-equality", [thm]);
val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs));
- val erhs = Envir.eta_contract (term_of rhs);
+ val erhs = Envir.eta_contract (Thm.term_of rhs);
val perm =
- var_perm (term_of elhs, erhs) andalso
- not (term_of elhs aconv erhs) andalso
- not (is_Var (term_of elhs));
- in (prems, term_of lhs, elhs, term_of rhs, perm) end;
+ var_perm (Thm.term_of elhs, erhs) andalso
+ not (Thm.term_of elhs aconv erhs) andalso
+ not (is_Var (Thm.term_of elhs));
+ in (prems, Thm.term_of lhs, elhs, Thm.term_of rhs, perm) end;
end;
@@ -534,7 +534,7 @@
if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
else
(*weak test for loops*)
- if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs)
+ if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (Thm.term_of elhs)
then mk_eq_True ctxt (thm, name)
else rrule_eq_True ctxt thm name lhs elhs rhs thm
end;
@@ -681,7 +681,7 @@
fun mk_simproc name lhss proc =
make_simproc {name = name, lhss = lhss, proc = fn _ => fn ctxt => fn ct =>
- proc ctxt (term_of ct), identifier = []};
+ proc ctxt (Thm.term_of ct), identifier = []};
(* FIXME avoid global thy and Logic.varify_global *)
fun simproc_global_i thy name = mk_simproc name o map (Thm.cterm_of thy o Logic.varify_global);
@@ -692,10 +692,10 @@
fun add_proc (proc as Proc {name, lhs, ...}) ctxt =
(cond_tracing ctxt (fn () =>
- print_term ctxt ("Adding simplification procedure " ^ quote name ^ " for") (term_of lhs));
+ print_term ctxt ("Adding simplification procedure " ^ quote name ^ " for") (Thm.term_of lhs));
ctxt |> map_simpset2
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
- (congs, Net.insert_term eq_proc (term_of lhs, proc) procs,
+ (congs, Net.insert_term eq_proc (Thm.term_of lhs, proc) procs,
mk_rews, termless, subgoal_tac, loop_tacs, solvers))
handle Net.INSERT =>
(cond_warning ctxt (fn () => "Ignoring duplicate simplification procedure " ^ quote name);
@@ -704,7 +704,7 @@
fun del_proc (proc as Proc {name, lhs, ...}) ctxt =
ctxt |> map_simpset2
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
- (congs, Net.delete_term eq_proc (term_of lhs, proc) procs,
+ (congs, Net.delete_term eq_proc (Thm.term_of lhs, proc) procs,
mk_rews, termless, subgoal_tac, loop_tacs, solvers))
handle Net.DELETE =>
(cond_warning ctxt (fn () => "Simplification procedure " ^ quote name ^ " not in simpset");
@@ -917,7 +917,7 @@
val Simpset ({rules, ...}, {congs, procs, termless, ...}) = simpset_of ctxt;
val eta_thm = Thm.eta_conversion t;
val eta_t' = Thm.rhs_of eta_thm;
- val eta_t = term_of eta_t';
+ val eta_t = Thm.term_of eta_t';
fun rew rrule =
let
val {thm, name, lhs, elhs, extra, fo, perm} = rrule
@@ -989,7 +989,7 @@
fun proc_rews [] = NONE
| proc_rews (Proc {name, proc, lhs, ...} :: ps) =
- if Pattern.matches (Proof_Context.theory_of ctxt) (term_of lhs, term_of t) then
+ if Pattern.matches (Proof_Context.theory_of ctxt) (Thm.term_of lhs, Thm.term_of t) then
(cond_tracing' ctxt simp_debug (fn () =>
print_term ctxt ("Trying procedure " ^ quote name ^ " on:") eta_t);
(case proc ctxt eta_t' of
@@ -1021,11 +1021,12 @@
fun congc prover ctxt maxt cong t =
let
val rthm = Thm.incr_indexes (maxt + 1) cong;
- val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm)));
+ val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (Thm.cprop_of rthm)));
val insts = Thm.match (rlhs, t)
(* Thm.match can raise Pattern.MATCH;
is handled when congc is called *)
- val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
+ val thm' =
+ Thm.instantiate insts (Thm.rename_boundvars (Thm.term_of rlhs) (Thm.term_of t) rthm);
val _ =
cond_tracing ctxt (fn () => print_thm ctxt "Applying congruence rule:" ("", thm'));
fun err (msg, thm) = (cond_tracing ctxt (fn () => print_thm ctxt msg ("", thm)); NONE);
@@ -1036,7 +1037,7 @@
(case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of
NONE => err ("Congruence proof failed. Should not have proved", thm2)
| SOME thm2' =>
- if op aconv (apply2 term_of (Thm.dest_equals (cprop_of thm2')))
+ if op aconv (apply2 Thm.term_of (Thm.dest_equals (Thm.cprop_of thm2')))
then NONE else SOME thm2'))
end;
@@ -1076,13 +1077,13 @@
and subc skel ctxt t0 =
let val Simpset (_, {congs, ...}) = simpset_of ctxt in
- (case term_of t0 of
+ (case Thm.term_of t0 of
Abs (a, T, _) =>
let
val (v, ctxt') = Variable.next_bound (a, T) ctxt;
val b = #1 (Term.dest_Free v);
val (v', t') = Thm.dest_abs (SOME b) t0;
- val b' = #1 (Term.dest_Free (term_of v'));
+ val b' = #1 (Term.dest_Free (Thm.term_of v'));
val _ =
if b <> b' then
warning ("Bad Simplifier context: renamed bound variable " ^
@@ -1155,11 +1156,11 @@
else nonmut_impc ct ctxt
and rules_of_prem prem ctxt =
- if maxidx_of_term (term_of prem) <> ~1
+ if maxidx_of_term (Thm.term_of prem) <> ~1
then
(cond_tracing ctxt (fn () =>
print_term ctxt "Cannot add premise as rewrite rule because it contains (type) unknowns:"
- (term_of prem));
+ (Thm.term_of prem));
(([], NONE), ctxt))
else
let val (asm, ctxt') = Thm.assume_hyps prem ctxt
@@ -1233,7 +1234,7 @@
| SOME eqn =>
let
val prem' = Thm.rhs_of eqn;
- val tprems = map term_of prems;
+ val tprems = map Thm.term_of prems;
val i = 1 + fold Integer.max (map (fn p =>
find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1;
val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt;
@@ -1294,18 +1295,18 @@
val ct = Thm.adjust_maxidx_cterm ~1 raw_ct;
val {maxidx, ...} = Thm.rep_cterm ct;
val _ =
- Theory.subthy (theory_of_cterm ct, thy) orelse
+ Theory.subthy (Thm.theory_of_cterm ct, thy) orelse
raise CTERM ("rewrite_cterm: bad background theory", [ct]);
val ctxt =
raw_ctxt
|> Context_Position.set_visible false
|> inc_simp_depth
- |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = term_of ct} ctxt);
+ |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = Thm.term_of ct} ctxt);
val _ =
cond_tracing ctxt (fn () =>
- print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (term_of ct));
+ print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (Thm.term_of ct));
in bottomc (mode, Option.map (Drule.flexflex_unique (SOME ctxt)) oo prover, maxidx) ctxt ct end;
val simple_prover =
@@ -1356,7 +1357,7 @@
| term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t)
| term_depth _ = 0;
-val lhs_of_thm = #1 o Logic.dest_equals o prop_of;
+val lhs_of_thm = #1 o Logic.dest_equals o Thm.prop_of;
(*folding should handle critical pairs! E.g. K == Inl(0), S == Inr(Inl(0))
Returns longest lhs first to avoid folding its subexpressions.*)
--- a/src/Pure/search.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/search.ML Wed Mar 04 19:53:18 2015 +0100
@@ -56,7 +56,7 @@
(*Predicate: Does the rule have fewer than n premises?*)
-fun has_fewer_prems n rule = (nprems_of rule < n);
+fun has_fewer_prems n rule = Thm.nprems_of rule < n;
(*Apply a tactic if subgoals remain, else do nothing.*)
val IF_UNSOLVED = COND (has_fewer_prems 1) all_tac;
@@ -67,16 +67,16 @@
(*Execute tac1, but only execute tac2 if there are at least as many subgoals
as before. This ensures that tac2 is only applied to an outcome of tac1.*)
fun (tac1 THEN_MAYBE tac2) st =
- (tac1 THEN COND (has_fewer_prems (nprems_of st)) all_tac tac2) st;
+ (tac1 THEN COND (has_fewer_prems (Thm.nprems_of st)) all_tac tac2) st;
fun (tac1 THEN_MAYBE' tac2) x = tac1 x THEN_MAYBE tac2 x;
(*Tactical to reduce the number of premises by 1.
If no subgoals then it must fail! *)
fun DEPTH_SOLVE_1 tac st = st |>
- (case nprems_of st of
- 0 => no_tac
- | n => DEPTH_FIRST (has_fewer_prems n) tac);
+ (case Thm.nprems_of st of
+ 0 => no_tac
+ | n => DEPTH_FIRST (has_fewer_prems n) tac);
(*Uses depth-first search to solve ALL subgoals*)
val DEPTH_SOLVE = DEPTH_FIRST (has_fewer_prems 1);
@@ -153,9 +153,9 @@
(fn()=> depth (bnd,inc) ((k,np,rgd,stq)::qs)))
else
- let val np' = nprems_of st
+ let val np' = Thm.nprems_of st
(*rgd' calculation assumes tactic operates on subgoal 1*)
- val rgd' = not (has_vars (hd (prems_of st)))
+ val rgd' = not (has_vars (hd (Thm.prems_of st)))
val k' = k+np'-np+1 (*difference in # of subgoals, +1*)
in if k'+np' >= bnd
then depth (bnd, Int.min(inc, k'+np'+1-bnd)) qs
--- a/src/Pure/tactic.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/tactic.ML Wed Mar 04 19:53:18 2015 +0100
@@ -206,7 +206,7 @@
fun filtr (limit, []) = []
| filtr (limit, th::ths) =
if limit=0 then []
- else if could(pb, concl_of th) then th :: filtr(limit-1, ths)
+ else if could(pb, Thm.concl_of th) then th :: filtr(limit-1, ths)
else filtr(limit,ths)
in filtr(limit,ths) end;
@@ -221,7 +221,7 @@
(case try Thm.major_prem_of th of
SOME prem => (inet, Net.insert_term (K false) (prem, kbrl) enet)
| NONE => error "insert_tagged_brl: elimination rule with no premises")
- else (Net.insert_term (K false) (concl_of th, kbrl) inet, enet);
+ else (Net.insert_term (K false) (Thm.concl_of th, kbrl) inet, enet);
(*delete one kbrl from the pair of nets*)
fun eq_kbrl ((_, (_, th)), (_, (_, th'))) = Thm.eq_thm_prop (th, th')
@@ -256,7 +256,7 @@
(*insert one tagged rl into the net*)
fun insert_krl (krl as (k,th)) =
- Net.insert_term (K false) (concl_of th, krl);
+ Net.insert_term (K false) (Thm.concl_of th, krl);
(*build a net of rules for resolution*)
fun build_net rls =
@@ -285,8 +285,8 @@
(*** For Natural Deduction using (bires_flg, rule) pairs ***)
(*The number of new subgoals produced by the brule*)
-fun subgoals_of_brl (true,rule) = nprems_of rule - 1
- | subgoals_of_brl (false,rule) = nprems_of rule;
+fun subgoals_of_brl (true, rule) = Thm.nprems_of rule - 1
+ | subgoals_of_brl (false, rule) = Thm.nprems_of rule;
(*Less-than test: for sorting to minimize number of new subgoals*)
fun lessb (brl1,brl2) = subgoals_of_brl brl1 < subgoals_of_brl brl2;
--- a/src/Pure/tactical.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/tactical.ML Wed Mar 04 19:53:18 2015 +0100
@@ -290,19 +290,19 @@
fun ALLGOALS tac st =
let fun doall 0 = all_tac
| doall n = tac(n) THEN doall(n-1)
- in doall(nprems_of st)st end;
+ in doall (Thm.nprems_of st) st end;
(*For n subgoals, performs tac(n) ORELSE ... ORELSE tac(1) *)
fun SOMEGOAL tac st =
let fun find 0 = no_tac
| find n = tac(n) ORELSE find(n-1)
- in find(nprems_of st)st end;
+ in find (Thm.nprems_of st) st end;
(*For n subgoals, performs tac(1) ORELSE ... ORELSE tac(n).
More appropriate than SOMEGOAL in some cases.*)
fun FIRSTGOAL tac st =
let fun find (i,n) = if i>n then no_tac else tac(i) ORELSE find (i+1,n)
- in find(1, nprems_of st)st end;
+ in find (1, Thm.nprems_of st) st end;
(*First subgoal only.*)
fun HEADGOAL tac = tac 1;
@@ -346,12 +346,12 @@
subgoal-based tactics this means subgoal i has been solved
altogether -- no new subgoals have emerged.*)
fun SOLVED' tac i st =
- tac i st |> Seq.filter (fn st' => nprems_of st' < nprems_of st);
+ tac i st |> Seq.filter (fn st' => Thm.nprems_of st' < Thm.nprems_of st);
(*Apply second tactic to all subgoals emerging from the first --
following usual convention for subgoal-based tactics.*)
fun (tac1 THEN_ALL_NEW tac2) i st =
- st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st) st'));
+ st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 i (i + Thm.nprems_of st' - Thm.nprems_of st) st'));
(*Repeatedly dig into any emerging subgoals.*)
fun REPEAT_ALL_NEW tac =
--- a/src/Pure/thm.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Pure/thm.ML Wed Mar 04 19:53:18 2015 +0100
@@ -9,26 +9,47 @@
signature BASIC_THM =
sig
+ type ctyp
+ type cterm
+ exception CTERM of string * cterm list
+ type thm
+ type conv = cterm -> thm
+ exception THM of string * int * thm list
+end;
+
+signature THM =
+sig
+ include BASIC_THM
+
(*certified types*)
- type ctyp
val rep_ctyp: ctyp -> {thy: theory, T: typ, maxidx: int, sorts: sort Ord_List.T}
val theory_of_ctyp: ctyp -> theory
val typ_of: ctyp -> typ
val ctyp_of: theory -> typ -> ctyp
+ val dest_ctyp: ctyp -> ctyp list
(*certified terms*)
- type cterm
- exception CTERM of string * cterm list
val rep_cterm: cterm -> {thy: theory, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}
val crep_cterm: cterm -> {thy: theory, t: term, T: ctyp, maxidx: int, sorts: sort Ord_List.T}
val theory_of_cterm: cterm -> theory
val term_of: cterm -> term
+ val ctyp_of_term: cterm -> ctyp
val cterm_of: theory -> term -> cterm
- val ctyp_of_term: cterm -> ctyp
+ val dest_comb: cterm -> cterm * cterm
+ val dest_fun: cterm -> cterm
+ val dest_arg: cterm -> cterm
+ val dest_fun2: cterm -> cterm
+ val dest_arg1: cterm -> cterm
+ val dest_abs: string option -> cterm -> cterm * cterm
+ val apply: cterm -> cterm -> cterm
+ val lambda_name: string * cterm -> cterm -> cterm
+ val lambda: cterm -> cterm -> cterm
+ val adjust_maxidx_cterm: int -> cterm -> cterm
+ val incr_indexes_cterm: int -> cterm -> cterm
+ val match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
+ val first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
(*theorems*)
- type thm
- type conv = cterm -> thm
val rep_thm: thm ->
{thy: theory,
tags: Properties.T,
@@ -45,42 +66,22 @@
hyps: cterm Ord_List.T,
tpairs: (cterm * cterm) list,
prop: cterm}
- exception THM of string * int * thm list
+ val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a
+ val terms_of_tpairs: (term * term) list -> term list
+ val full_prop_of: thm -> term
val theory_of_thm: thm -> theory
+ val maxidx_of: thm -> int
+ val maxidx_thm: thm -> int -> int
+ val hyps_of: thm -> term list
val prop_of: thm -> term
+ val tpairs_of: thm -> (term * term) list
val concl_of: thm -> term
val prems_of: thm -> term list
val nprems_of: thm -> int
+ val no_prems: thm -> bool
+ val major_prem_of: thm -> term
val cprop_of: thm -> cterm
val cprem_of: thm -> int -> cterm
-end;
-
-signature THM =
-sig
- include BASIC_THM
- val dest_ctyp: ctyp -> ctyp list
- val dest_comb: cterm -> cterm * cterm
- val dest_fun: cterm -> cterm
- val dest_arg: cterm -> cterm
- val dest_fun2: cterm -> cterm
- val dest_arg1: cterm -> cterm
- val dest_abs: string option -> cterm -> cterm * cterm
- val apply: cterm -> cterm -> cterm
- val lambda_name: string * cterm -> cterm -> cterm
- val lambda: cterm -> cterm -> cterm
- val adjust_maxidx_cterm: int -> cterm -> cterm
- val incr_indexes_cterm: int -> cterm -> cterm
- val match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
- val first_order_match: cterm * cterm -> (ctyp * ctyp) list * (cterm * cterm) list
- val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a
- val terms_of_tpairs: (term * term) list -> term list
- val full_prop_of: thm -> term
- val maxidx_of: thm -> int
- val maxidx_thm: thm -> int -> int
- val hyps_of: thm -> term list
- val tpairs_of: thm -> (term * term) list
- val no_prems: thm -> bool
- val major_prem_of: thm -> term
val transfer: theory -> thm -> thm
val weaken: cterm -> thm -> thm
val weaken_sorts: sort list -> cterm -> cterm
--- a/src/Sequents/modal.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Sequents/modal.ML Wed Mar 04 19:53:18 2015 +0100
@@ -69,12 +69,12 @@
fun UPTOGOAL n tf = let fun tac i = if i<n then all_tac
else tf(i) THEN tac(i-1)
- in fn st => tac (nprems_of st) st end;
+ in fn st => tac (Thm.nprems_of st) st end;
(* Depth first search bounded by d *)
fun solven_tac ctxt d n st = st |>
(if d < 0 then no_tac
- else if nprems_of st = 0 then all_tac
+ else if Thm.nprems_of st = 0 then all_tac
else (DETERM(fres_safe_tac ctxt n) THEN UPTOGOAL n (solven_tac ctxt d)) ORELSE
((fres_unsafe_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt d)) APPEND
(fres_bound_tac ctxt n THEN UPTOGOAL n (solven_tac ctxt (d - 1)))));
--- a/src/Sequents/prover.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Sequents/prover.ML Wed Mar 04 19:53:18 2015 +0100
@@ -157,7 +157,7 @@
(*Predicate: does the rule have n premises? *)
-fun has_prems n rule = (nprems_of rule = n);
+fun has_prems n rule = (Thm.nprems_of rule = n);
(*Continuation-style tactical for resolution.
The list of rules is partitioned into 0, 1, 2 premises.
--- a/src/Sequents/simpdata.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Sequents/simpdata.ML Wed Mar 04 19:53:18 2015 +0100
@@ -12,7 +12,7 @@
(*Make atomic rewrite rules*)
fun atomize r =
- case concl_of r of
+ case Thm.concl_of r of
Const(@{const_name Trueprop},_) $ Abs(_,_,a) $ Abs(_,_,c) =>
(case (Cla.forms_of_seq a, Cla.forms_of_seq c) of
([], [p]) =>
@@ -28,7 +28,7 @@
| _ => [r];
(*Make meta-equalities.*)
-fun mk_meta_eq ctxt th = case concl_of th of
+fun mk_meta_eq ctxt th = case Thm.concl_of th of
Const(@{const_name Pure.eq},_)$_$_ => th
| Const(@{const_name Trueprop},_) $ Abs(_,_,a) $ Abs(_,_,c) =>
(case (Cla.forms_of_seq a, Cla.forms_of_seq c) of
--- a/src/Tools/Code/code_preproc.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/Code/code_preproc.ML Wed Mar 04 19:53:18 2015 +0100
@@ -149,16 +149,16 @@
fun conversion sandwich conv ctxt ct =
let
val (postproc, ct') = sandwich ctxt ct;
- in postproc (conv ctxt (term_of ct') ct') end;
+ in postproc (conv ctxt (Thm.term_of ct') ct') end;
fun evaluation sandwich lift_postproc eval ctxt t =
let
val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
val (postproc, ct') = sandwich ctxt (cert t);
in
- term_of ct'
+ Thm.term_of ct'
|> eval ctxt
- |> lift_postproc (term_of o Thm.rhs_of o postproc o Thm.reflexive o cert)
+ |> lift_postproc (Thm.term_of o Thm.rhs_of o postproc o Thm.reflexive o cert)
end;
end;
@@ -168,8 +168,8 @@
fun normalized_tfrees_sandwich ctxt ct =
let
- val cert = cterm_of (Proof_Context.theory_of ctxt);
- val t = term_of ct;
+ val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
+ val t = Thm.term_of ct;
val vs_original = fold_term_types (K (fold_atyps (insert (eq_fst op =)
o dest_TFree))) t [];
val vs_normalized = Name.invent_names Name.context Name.aT (map snd vs_original);
--- a/src/Tools/IsaPlanner/isand.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/IsaPlanner/isand.ML Wed Mar 04 19:53:18 2015 +0100
@@ -142,7 +142,7 @@
let
val cert = Thm.cterm_of (Thm.theory_of_thm th);
- val t = prop_of th;
+ val t = Thm.prop_of th;
val prems = Logic.strip_imp_prems t;
val aprems = map (Thm.trivial o cert) prems;
--- a/src/Tools/atomize_elim.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/atomize_elim.ML Wed Mar 04 19:53:18 2015 +0100
@@ -39,7 +39,7 @@
val pi' = 0 :: map (Integer.add 1) pi
val fwd = Thm.trivial ct
|> Drule.rearrange_prems pi'
- val back = Thm.trivial (snd (Thm.dest_implies (cprop_of fwd)))
+ val back = Thm.trivial (snd (Thm.dest_implies (Thm.cprop_of fwd)))
|> Drule.rearrange_prems (invert_perm pi')
in Thm.equal_intr fwd back end
@@ -73,7 +73,7 @@
(* move unrelated assumptions out of the quantification *)
fun movea_conv ctxt ct =
let
- val _ $ Abs (_, _, b) = term_of ct
+ val _ $ Abs (_, _, b) = Thm.term_of ct
val idxs = fold_index (fn (i, t) => not (Term.is_dependent t) ? cons i)
(Logic.strip_imp_prems b) []
|> rev
@@ -92,7 +92,7 @@
else_conv (movea_conv ctxt)
fun ms_conv ctxt ct =
- if is_forall_thesis (term_of ct)
+ if is_forall_thesis (Thm.term_of ct)
then moveq_conv ctxt ct
else (implies_concl_conv (ms_conv ctxt)
else_conv
@@ -116,8 +116,8 @@
val quantify_thesis =
if is_Bound thesis then all_tac
else let
- val cthesis = cterm_of thy thesis
- val rule = instantiate' [SOME (ctyp_of_term cthesis)] [NONE, SOME cthesis]
+ val cthesis = Thm.cterm_of thy thesis
+ val rule = instantiate' [SOME (Thm.ctyp_of_term cthesis)] [NONE, SOME cthesis]
@{thm meta_spec}
in
compose_tac ctxt (false, rule, 1) i
--- a/src/Tools/coherent.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/coherent.ML Wed Mar 04 19:53:18 2015 +0100
@@ -39,8 +39,8 @@
val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1);
fun rulify_elim_conv ctxt ct =
- if is_atomic (Logic.strip_imp_concl (term_of ct)) then Conv.all_conv ct
- else Conv.concl_conv (length (Logic.strip_imp_prems (term_of ct)))
+ if is_atomic (Logic.strip_imp_concl (Thm.term_of ct)) then Conv.all_conv ct
+ else Conv.concl_conv (length (Logic.strip_imp_prems (Thm.term_of ct)))
(Conv.rewr_conv (Thm.symmetric Data.atomize_elimL) then_conv
Raw_Simplifier.rewrite ctxt true (map Thm.symmetric
[Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct
@@ -66,7 +66,7 @@
fun mk_rule ctxt th =
let
val th' = rulify_elim ctxt th;
- val (prems, cases) = dest_elim (prop_of th')
+ val (prems, cases) = dest_elim (Thm.prop_of th')
in (th', prems, cases) end;
fun mk_dom ts = fold (fn t =>
@@ -187,7 +187,7 @@
(cert (Var (ixn, Envir.subst_type tye T)), cert t)) (Vartab.dest env) @
map (fn (ixnT, t) => (cert (Var ixnT), cert t)) insts) th)
(map (nth asms) is);
- val (_, cases) = dest_elim (prop_of th');
+ val (_, cases) = dest_elim (Thm.prop_of th');
in
(case (cases, prfs) of
([([], [_])], []) => th'
@@ -195,7 +195,7 @@
| _ =>
Drule.implies_elim_list
(Thm.instantiate (Thm.match
- (Drule.strip_imp_concl (cprop_of th'), goal)) th')
+ (Drule.strip_imp_concl (Thm.cprop_of th'), goal)) th')
(map (thm_of_case_prf ctxt goal asms) (prfs ~~ cases)))
end
@@ -219,12 +219,12 @@
SUBPROOF (fn {prems = prems', concl, context = ctxt'', ...} =>
let
val xs =
- map (term_of o #2) params @
+ map (Thm.term_of o #2) params @
map (fn (_, s) => Free (s, the (Variable.default_type ctxt'' s)))
(rev (Variable.dest_fixes ctxt'')) (* FIXME !? *)
in
(case
- valid ctxt'' (map (mk_rule ctxt'') (prems' @ prems @ rules)) (term_of concl)
+ valid ctxt'' (map (mk_rule ctxt'') (prems' @ prems @ rules)) (Thm.term_of concl)
(mk_dom xs) Net.empty 0 0 of
NONE => no_tac
| SOME prf => resolve_tac ctxt'' [thm_of_cl_prf ctxt'' concl [] prf] 1)
--- a/src/Tools/induct.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/induct.ML Wed Mar 04 19:53:18 2015 +0100
@@ -156,13 +156,13 @@
end;
fun mk_swap_rrule ctxt ct =
- (case find_eq ctxt (term_of ct) of
+ (case find_eq ctxt (Thm.term_of ct) of
NONE => NONE
| SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct));
val rearrange_eqs_simproc =
Simplifier.simproc_global Pure.thy "rearrange_eqs" ["Pure.all t"]
- (fn ctxt => fn t => mk_swap_rrule ctxt (cterm_of (Proof_Context.theory_of ctxt) t));
+ (fn ctxt => fn t => mk_swap_rrule ctxt (Thm.cterm_of (Proof_Context.theory_of ctxt) t));
(* rotate k premises to the left by j, skipping over first j premises *)
@@ -180,8 +180,8 @@
(* rulify operators around definition *)
fun rulify_defs_conv ctxt ct =
- if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) andalso
- not (is_some (Induct_Args.dest_def (drop_judgment ctxt (term_of ct))))
+ if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) andalso
+ not (is_some (Induct_Args.dest_def (drop_judgment ctxt (Thm.term_of ct))))
then
(Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv
Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt))
@@ -440,7 +440,7 @@
Simplifier.full_rewrite (put_simpset (#4 (get_local ctxt)) ctxt);
fun simplify_conv ctxt ct =
- if exists_subterm (is_some o Induct_Args.dest_def) (term_of ct) then
+ if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) then
(Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct
else Conv.all_conv ct;
--- a/src/Tools/induction.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/induction.ML Wed Mar 04 19:53:18 2015 +0100
@@ -20,7 +20,7 @@
if not(forall (null o #2 o #1) cases) then arg
else
let
- val (prems, concl) = Logic.strip_horn (prop_of th);
+ val (prems, concl) = Logic.strip_horn (Thm.prop_of th);
val prems' = drop consumes prems;
val ps = preds_of concl;
--- a/src/Tools/misc_legacy.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/misc_legacy.ML Wed Mar 04 19:53:18 2015 +0100
@@ -193,7 +193,7 @@
and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
- val emBs = map (cterm o embed) (prems_of st')
+ val emBs = map (cterm o embed) (Thm.prems_of st')
val Cth = implies_elim_list st' (map (elim o Thm.assume) emBs)
in (*restore the unknowns to the hypotheses*)
free_instantiate (map swap_ctpair insts @
@@ -205,7 +205,7 @@
(*function to replace the current subgoal*)
fun next st =
Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false}
- (false, relift st, nprems_of st) gno state
+ (false, relift st, Thm.nprems_of st) gno state
in Seq.maps next (tacf subprems (Thm.trivial (cterm concl))) end;
fun print_vars_terms ctxt n thm =
@@ -265,8 +265,8 @@
let fun newName (ix,_) = (ix, gensym (string_of_indexname ix))
val alist = map newName vars
fun mk_inst (v,T) =
- (cterm_of thy (Var(v,T)),
- cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))
+ (Thm.cterm_of thy (Var(v,T)),
+ Thm.cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))
val insts = map mk_inst vars
fun thaw i th' = (*i is non-negative increment for Var indexes*)
th' |> forall_intr_list (map #2 insts)
@@ -289,8 +289,8 @@
val (alist, _) =
fold_rev newName vars ([], Thm.fold_terms Term.add_free_names fth [])
fun mk_inst (v, T) =
- (cterm_of thy (Var(v,T)),
- cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))
+ (Thm.cterm_of thy (Var(v,T)),
+ Thm.cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T)))
val insts = map mk_inst vars
fun thaw th' =
th' |> forall_intr_list (map #2 insts)
--- a/src/Tools/nbe.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/nbe.ML Wed Mar 04 19:53:18 2015 +0100
@@ -589,7 +589,7 @@
val (_, raw_oracle) = Context.>>> (Context.map_theory_result
(Thm.add_oracle (@{binding normalization_by_evaluation},
fn (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct) =>
- mk_equals ctxt ct (eval_term nbe_program_idx_tab ctxt (term_of ct) vs_ty_t deps))));
+ mk_equals ctxt ct (eval_term nbe_program_idx_tab ctxt (Thm.term_of ct) vs_ty_t deps))));
fun oracle nbe_program_idx_tab ctxt vs_ty_t deps ct =
raw_oracle (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct);
--- a/src/Tools/quickcheck.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/quickcheck.ML Wed Mar 04 19:53:18 2015 +0100
@@ -344,7 +344,7 @@
fun strip (Const (@{const_name Pure.all}, _) $ Abs (_, _, t)) = strip t
| strip t = t;
val {goal = st, ...} = Proof.raw_goal state;
- val (gi, frees) = Logic.goal_params (prop_of st) i;
+ val (gi, frees) = Logic.goal_params (Thm.prop_of st) i;
val some_locale = Named_Target.bottom_locale_of lthy;
val assms =
if Config.get lthy no_assms then []
--- a/src/Tools/try.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/Tools/try.ML Wed Mar 04 19:53:18 2015 +0100
@@ -35,7 +35,7 @@
| serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
| serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
-val subgoal_count = Logic.count_prems o prop_of o #goal o Proof.goal
+val subgoal_count = Logic.count_prems o Thm.prop_of o #goal o Proof.goal
(* configuration *)
--- a/src/ZF/Tools/datatype_package.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/Tools/datatype_package.ML Wed Mar 04 19:53:18 2015 +0100
@@ -345,7 +345,7 @@
con1(x)=con1(y) ==> x=y, con1(x)=con1(y) <-> x=y, etc. *)
fun mk_free s =
let
- val thy = theory_of_thm elim;
+ val thy = Thm.theory_of_thm elim;
val ctxt = Proof_Context.init_global thy;
in
Goal.prove_global thy [] [] (Syntax.read_prop ctxt s)
--- a/src/ZF/Tools/induct_tacs.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/Tools/induct_tacs.ML Wed Mar 04 19:53:18 2015 +0100
@@ -92,12 +92,12 @@
let
val thy = Proof_Context.theory_of ctxt
val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i state
- val tn = find_tname ctxt' var (map term_of asms)
+ val tn = find_tname ctxt' var (map Thm.term_of asms)
val rule =
if exh then #exhaustion (datatype_info thy tn)
else #induct (datatype_info thy tn)
val (Const(@{const_name mem},_) $ Var(ixn,_) $ _) =
- (case prems_of rule of
+ (case Thm.prems_of rule of
[] => error "induction is not available for this datatype"
| major::_ => FOLogic.dest_Trueprop major)
in
@@ -125,7 +125,7 @@
map (head_of o const_of o FOLogic.dest_Trueprop o Thm.prop_of) case_eqns;
val Const (@{const_name mem}, _) $ _ $ data =
- FOLogic.dest_Trueprop (hd (prems_of elim));
+ FOLogic.dest_Trueprop (hd (Thm.prems_of elim));
val Const(big_rec_name, _) = head_of data;
--- a/src/ZF/Tools/inductive_package.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/Tools/inductive_package.ML Wed Mar 04 19:53:18 2015 +0100
@@ -494,13 +494,13 @@
The name "x.1" comes from the "RS spec" !*)
val inst =
case elem_frees of [_] => I
- | _ => Drule.instantiate_normalize ([], [(cterm_of thy (Var(("x",1), Ind_Syntax.iT)),
- cterm_of thy elem_tuple)]);
+ | _ => Drule.instantiate_normalize ([], [(Thm.cterm_of thy (Var(("x",1), Ind_Syntax.iT)),
+ Thm.cterm_of thy elem_tuple)]);
(*strip quantifier and the implication*)
val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
- val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
+ val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = Thm.concl_of induct0
val induct =
CP.split_rule_var (Proof_Context.init_global thy)
--- a/src/ZF/Tools/primrec_package.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/Tools/primrec_package.ML Wed Mar 04 19:53:18 2015 +0100
@@ -103,7 +103,7 @@
| use_fabs t = t
val cnames = map (#1 o dest_Const) constructors
- and recursor_pairs = map (dest_eqn o concl_of) rec_rewrites
+ and recursor_pairs = map (dest_eqn o Thm.concl_of) rec_rewrites
fun absterm (Free x, body) = absfree x body
| absterm (t, body) = Abs("rec", Ind_Syntax.iT, abstract_over (t, body))
--- a/src/ZF/arith_data.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/arith_data.ML Wed Mar 04 19:53:18 2015 +0100
@@ -63,7 +63,7 @@
fun is_eq_thm th =
can FOLogic.dest_eq (FOLogic.dest_Trueprop (Thm.prop_of th));
-fun add_chyps chyps ct = Drule.list_implies (map cprop_of chyps, ct);
+fun add_chyps chyps ct = Drule.list_implies (map Thm.cprop_of chyps, ct);
fun prove_conv name tacs ctxt prems (t,u) =
if t aconv u then NONE
--- a/src/ZF/simpdata.ML Tue Mar 03 19:08:04 2015 +0100
+++ b/src/ZF/simpdata.ML Wed Mar 04 19:53:18 2015 +0100
@@ -18,7 +18,7 @@
SOME rls => maps (atomize (conn_pairs, mem_pairs)) ([th] RL rls)
| NONE => [th])
| _ => [th]
- in case concl_of th of
+ in case Thm.concl_of th of
Const(@{const_name Trueprop},_) $ P =>
(case P of
Const(@{const_name mem},_) $ a $ b => tryrules mem_pairs b