# HG changeset patch # User wenzelm # Date 1425653936 -3600 # Node ID 291934bac95e3aeca5834bd2660606fa0039e854 # Parent 92d7d8e4f1bf7c7b0987e1bede5921a4a4d438bd Thm.cterm_of and Thm.ctyp_of operate on local context; diff -r 92d7d8e4f1bf -r 291934bac95e NEWS --- a/NEWS Fri Mar 06 14:01:08 2015 +0100 +++ b/NEWS Fri Mar 06 15:58:56 2015 +0100 @@ -263,6 +263,11 @@ *** ML *** +* The main operations to certify logical entities are Thm.ctyp_of and +Thm.cterm_of with a local context; old-style global theory variants are +available as Thm.global_ctyp_of and Thm.global_cterm_of. +INCOMPATIBILITY. + * Elementary operations in module Thm are no longer pervasive. INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of, Thm.term_of etc. diff -r 92d7d8e4f1bf -r 291934bac95e src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Doc/Implementation/Logic.thy Fri Mar 06 15:58:56 2015 +0100 @@ -639,8 +639,8 @@ \begin{mldecls} @{index_ML_type ctyp} \\ @{index_ML_type cterm} \\ - @{index_ML Thm.ctyp_of: "theory -> typ -> ctyp"} \\ - @{index_ML Thm.cterm_of: "theory -> term -> cterm"} \\ + @{index_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ + @{index_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{index_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{index_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{index_ML Thm.all: "cterm -> cterm -> cterm"} \\ @@ -699,10 +699,11 @@ are located in the @{ML_structure Thm} module, even though theorems are not yet involved at that stage. - \item @{ML Thm.ctyp_of}~@{text "thy \"} and @{ML - Thm.cterm_of}~@{text "thy t"} explicitly checks types and terms, + \item @{ML Thm.ctyp_of}~@{text "ctxt \"} and @{ML + Thm.cterm_of}~@{text "ctxt t"} explicitly check types and terms, respectively. This also involves some basic normalizations, such - expansion of type and term abbreviations from the theory context. + expansion of type and term abbreviations from the underlying + theory context. Full re-certification is relatively slow and should be avoided in tight reasoning loops. diff -r 92d7d8e4f1bf -r 291934bac95e src/FOLP/simp.ML --- a/src/FOLP/simp.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/FOLP/simp.ML Fri Mar 06 15:58:56 2015 +0100 @@ -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 Thm.cterm_of(Thm.theory_of_thm thm)) As; + val thms = map (Thm.trivial o Thm.global_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; @@ -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 = 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)) + let val ca = Thm.global_cterm_of sg va + and cx = Thm.global_cterm_of sg (eta_Var(("X"^si,0),T)) + val cb = Thm.global_cterm_of sg vb + and cy = Thm.global_cterm_of sg (eta_Var(("Y"^si,0),T)) + val cP = Thm.global_cterm_of sg P + and cp = Thm.global_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) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Fri Mar 06 15:58:56 2015 +0100 @@ -3582,7 +3582,7 @@ val normalize = eval o Envir.beta_norm o Envir.eta_long []; -in Thm.cterm_of thy (Logic.mk_equals (t, normalize t)) end +in Thm.global_cterm_of thy (Logic.mk_equals (t, normalize t)) end *} ML {* @@ -3641,16 +3641,16 @@ |> HOLogic.dest_list val p = prec |> HOLogic.mk_number @{typ nat} - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt in case taylor of NONE => let val n = vs |> length |> HOLogic.mk_number @{typ nat} - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> HOLogic.mk_list @{typ nat} - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt in (rtac (Thm.instantiate ([], [(@{cpat "?n::nat"}, n), (@{cpat "?prec::nat"}, p), @@ -3665,9 +3665,9 @@ else let val t = t |> HOLogic.mk_number @{typ nat} - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> hd - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt in rtac (Thm.instantiate ([], [(@{cpat "?s::nat"}, s), (@{cpat "?t::nat"}, t), diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Fri Mar 06 15:58:56 2015 +0100 @@ -2649,7 +2649,7 @@ val vs = map_index swap fs; val ps = map_index swap bs; val t' = (term_of_fm ps vs o @{code pa} o fm_of_term ps vs) t; - in Thm.cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end + in Thm.global_cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end end; *} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/Ferrack.thy Fri Mar 06 15:58:56 2015 +0100 @@ -1999,7 +1999,7 @@ let val vs = Term.add_frees t []; val t' = (term_of_fm vs o @{code linrqe} o fm_of_term vs) t; - in (Proof_Context.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end + in (Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end end; *} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Fri Mar 06 15:58:56 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 (Thm.cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end + in (Thm.global_cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end end; *} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Mar 06 15:58:56 2015 +0100 @@ -4155,7 +4155,7 @@ in fn (ctxt, alternative, ty, ps, ct) => - Proof_Context.cterm_of ctxt + Thm.cterm_of ctxt (frpar_procedure alternative ty ps (Thm.term_of ct)) end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/approximation.ML --- a/src/HOL/Decision_Procs/approximation.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/approximation.ML Fri Mar 06 15:58:56 2015 +0100 @@ -92,7 +92,7 @@ in t end fun apply_tactic ctxt term tactic = - Proof_Context.cterm_of ctxt term + Thm.cterm_of ctxt term |> Goal.init |> SINGLE tactic |> the |> Thm.prems_of |> hd @@ -120,7 +120,7 @@ |> foldr1 HOLogic.mk_conj)) fun approx_arith prec ctxt t = realify t - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt |> Reification.conv ctxt form_equations |> Thm.prop_of |> Logic.dest_equals |> snd diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/approximation_generator.ML --- a/src/HOL/Decision_Procs/approximation_generator.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/approximation_generator.ML Fri Mar 06 15:58:56 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 = Thm.cterm_of thy r |> conv |> Thm.prop_of |> Logic.dest_equals |> snd +fun conv_term thy conv r = Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/commutative_ring_tac.ML --- a/src/HOL/Decision_Procs/commutative_ring_tac.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Fri Mar 06 15:58:56 2015 +0100 @@ -78,10 +78,10 @@ let val thy = Proof_Context.theory_of ctxt; val fs = Misc_Legacy.term_frees eq; - 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; + val cvs = Thm.global_cterm_of thy (HOLogic.mk_list T fs); + val clhs = Thm.global_cterm_of thy (reif_polex T fs lhs); + val crhs = Thm.global_cterm_of thy (reif_polex T fs rhs); + val ca = Thm.global_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"; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/cooper_tac.ML --- a/src/HOL/Decision_Procs/cooper_tac.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/cooper_tac.ML Fri Mar 06 15:58:56 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 = Thm.cterm_of thy (HOLogic.mk_Trueprop t) + val ct = Thm.global_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, @@ -86,7 +86,7 @@ (case Thm.prop_of pre_thm of Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ => let - val pth = linzqe_oracle (Thm.cterm_of thy (Envir.eta_long [] t1)) + val pth = linzqe_oracle (Thm.global_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)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/ferrack_tac.ML --- a/src/HOL/Decision_Procs/ferrack_tac.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/ferrack_tac.ML Fri Mar 06 15:58:56 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 = Thm.cterm_of thy (HOLogic.mk_Trueprop t) + val ct = Thm.global_cterm_of thy (HOLogic.mk_Trueprop t) (* Theorem for the nat --> int transformation *) val pre_thm = Seq.hd (EVERY [simp_tac simpset0 1, diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Decision_Procs/mir_tac.ML Fri Mar 06 15:58:56 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 = Thm.cterm_of thy (HOLogic.mk_Trueprop t) + val ct = Thm.global_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, @@ -116,8 +116,8 @@ 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, Thm.cterm_of thy (Envir.eta_long [] t1)) - else mirfr_oracle (true, Thm.cterm_of thy (Envir.eta_long [] t1)) + then mirfr_oracle (false, Thm.global_cterm_of thy (Envir.eta_long [] t1)) + else mirfr_oracle (true, Thm.global_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)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/HOL.thy --- a/src/HOL/HOL.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/HOL.thy Fri Mar 06 15:58:56 2015 +0100 @@ -1194,13 +1194,13 @@ let val (f_Let_unfold, x_Let_unfold) = let val [(_ $ (f $ x) $ _)] = Thm.prems_of @{thm Let_unfold} - in (Thm.cterm_of @{theory} f, Thm.cterm_of @{theory} x) end + in (Thm.global_cterm_of @{theory} f, Thm.global_cterm_of @{theory} x) end val (f_Let_folded, x_Let_folded) = let val [(_ $ (f $ x) $ _)] = Thm.prems_of @{thm Let_folded} - in (Thm.cterm_of @{theory} f, Thm.cterm_of @{theory} x) end; + in (Thm.global_cterm_of @{theory} f, Thm.global_cterm_of @{theory} x) end; val g_Let_folded = let val [(_ $ _ $ (g $ _))] = Thm.prems_of @{thm Let_folded} - in Thm.cterm_of @{theory} g end; + in Thm.global_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,9 +1222,9 @@ else let val n = case f of (Abs (x, _, _)) => x | _ => "x"; - val cx = Thm.cterm_of thy x; + val cx = Thm.global_cterm_of thy x; val xT = Thm.typ_of_cterm cx; - val cf = Thm.cterm_of thy f; + val cf = Thm.global_cterm_of thy f; val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx); val (_ $ _ $ g) = Thm.prop_of fx_g; val g' = abstract_over (x,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 (Thm.cterm_of thy g'x)); + val g_g'x = Thm.symmetric (Thm.beta_conversion false (Thm.global_cterm_of thy g'x)); val rl = cterm_instantiate - [(f_Let_folded, Thm.cterm_of thy f), (x_Let_folded, cx), - (g_Let_folded, Thm.cterm_of thy abs_g')] + [(f_Let_folded, Thm.global_cterm_of thy f), (x_Let_folded, cx), + (g_Let_folded, Thm.global_cterm_of thy abs_g')] @{thm Let_folded}; in SOME (rl OF [Thm.transitive fx_g g_g'x]) end) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/HOLCF/Tools/Domain/domain_constructors.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/HOLCF/Tools/Domain/domain_constructors.ML Fri Mar 06 15:58:56 2015 +0100 @@ -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 = Thm.ctyp_of thy (snd (cons2typ 1 spec')) + val ct = Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/HOLCF/Tools/cont_proc.ML --- a/src/HOL/HOLCF/Tools/cont_proc.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/HOLCF/Tools/cont_proc.ML Fri Mar 06 15:58:56 2015 +0100 @@ -122,7 +122,7 @@ fun solve_cont ctxt t = let val thy = Proof_Context.theory_of ctxt - val tr = instantiate' [] [SOME (Thm.cterm_of thy t)] @{thm Eq_TrueI} + val tr = instantiate' [] [SOME (Thm.global_cterm_of thy t)] @{thm Eq_TrueI} in Option.map fst (Seq.pull (cont_tac ctxt 1 tr)) end in fun cont_proc thy = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/HOLCF/Tools/fixrec.ML --- a/src/HOL/HOLCF/Tools/fixrec.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/HOLCF/Tools/fixrec.ML Fri Mar 06 15:58:56 2015 +0100 @@ -153,7 +153,7 @@ val P = Var (("P", 0), map Term.fastype_of lhss ---> HOLogic.boolT) val predicate = lambda_tuple lhss (list_comb (P, lhss)) val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm]) - |> Drule.instantiate' [] [SOME (Thm.cterm_of thy predicate)] + |> Drule.instantiate' [] [SOME (Thm.global_cterm_of thy predicate)] |> Local_Defs.unfold lthy @{thms split_paired_all split_conv split_strict} val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) |> Local_Defs.unfold lthy @{thms split_conv} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Import/import_rule.ML --- a/src/HOL/Import/import_rule.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Import/import_rule.ML Fri Mar 06 15:58:56 2015 +0100 @@ -163,7 +163,7 @@ 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 = Thm.ctyp_of thy ty + fun inst ty = Thm.global_ctyp_of thy ty in Thm.instantiate ((map inst tvars ~~ map inst tfrees), []) thm end @@ -206,8 +206,8 @@ val [nty, oty] = Thm.dest_ctyp (Thm.ctyp_of_cterm rept) in Drule.instantiate' [SOME nty, SOME oty] [SOME rept, SOME abst, SOME P, - 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} + SOME (Thm.global_cterm_of thy (Free ("a", Thm.typ_of nty))), + SOME (Thm.global_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 = @@ -233,8 +233,8 @@ val typedef_th = Drule.instantiate' [SOME nty, SOME oty] - [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)))] + [SOME rept, SOME abst, SOME cP, SOME (Thm.global_cterm_of thy' (Free ("a", aty))), + SOME (Thm.global_cterm_of thy' (Free ("r", Thm.typ_of ctT)))] @{thm typedef_hol2hollight} val th4 = typedef_th OF [#type_definition (#2 typedef_info)] in @@ -266,8 +266,8 @@ 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 => (Thm.ctyp_of thy (TVar iS), cty) - | NONE => (Thm.ctyp_of thy (TVar iS), Thm.ctyp_of thy (TFree bef)) + SOME cty => (Thm.global_ctyp_of thy (TVar iS), cty) + | NONE => (Thm.global_ctyp_of thy (TVar iS), Thm.global_ctyp_of thy (TFree bef)) )) tys_before tys_after in Thm.instantiate (tyinst,[]) th1 @@ -333,8 +333,8 @@ 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 (Thm.ctyp_of thy) vs - val ctvs = map (Thm.ctyp_of thy) (map TVar tvs) + val cvs = map (Thm.global_ctyp_of thy) vs + val ctvs = map (Thm.global_ctyp_of thy) (map TVar tvs) val thm' = Thm.instantiate ((ctvs ~~ cvs), []) thm in snd (Global_Theory.add_thm ((binding, thm'), []) thy) @@ -411,14 +411,14 @@ end | process (thy, state) (#"Y", [name, _, _]) = setth (mtydef name thy) (thy, state) | process (thy, state) (#"t", [n]) = - setty (Thm.ctyp_of thy (TFree ("'" ^ (transl_qm n), @{sort type}))) (thy, state) + setty (Thm.global_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 => Thm.ctyp_of thy (Type (gettyname n thy, map Thm.typ_of tys))) |-> setty + (fn tys => Thm.global_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 => Thm.cterm_of thy (Free (transl_dot n, Thm.typ_of ty))) |-> settm + getty ty (thy, state) |>> (fn ty => Thm.global_cterm_of thy (Free (transl_dot n, Thm.typ_of ty))) |-> settm | process (thy, state) (#"c", [n, ty]) = - getty ty (thy, state) |>> (fn ty => Thm.cterm_of thy (Const (getconstname n thy, Thm.typ_of ty))) |-> settm + getty ty (thy, state) |>> (fn ty => Thm.global_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]) = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Code_Abstract_Nat.thy --- a/src/HOL/Library/Code_Abstract_Nat.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Code_Abstract_Nat.thy Fri Mar 06 15:58:56 2015 +0100 @@ -59,7 +59,7 @@ 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 = Thm.cterm_of thy (Var ((vname, 0), HOLogic.natT)); + val cv = Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Countable.thy --- a/src/HOL/Library/Countable.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Countable.thy Fri Mar 06 15:58:56 2015 +0100 @@ -181,7 +181,7 @@ val induct_thm = the (AList.lookup (op =) alist pred_name) val vars = rev (Term.add_vars (Thm.prop_of induct_thm) []) val thy = Proof_Context.theory_of ctxt - val insts = vars |> map (fn (_, T) => try (Thm.cterm_of thy) + val insts = vars |> map (fn (_, T) => try (Thm.global_cterm_of thy) (Const (@{const_name Countable.finite_item}, T))) val induct_thm' = Drule.instantiate' [] insts induct_thm val rules = @{thms finite_item.intros} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_smt_normalize.ML --- a/src/HOL/Library/Old_SMT/old_smt_normalize.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_smt_normalize.ML Fri Mar 06 15:58:56 2015 +0100 @@ -30,7 +30,7 @@ local val (cpfalse, cfalse) = - `Old_SMT_Utils.mk_cprop (Thm.cterm_of @{theory} @{const False}) + `Old_SMT_Utils.mk_cprop (Thm.global_cterm_of @{theory} @{const False}) fun inst f ct thm = let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) @@ -203,7 +203,7 @@ Old_SMT_Utils.mk_const_pat @{theory} @{const_name pat} Old_SMT_Utils.destT1 fun mk_pat ct = Thm.apply (Old_SMT_Utils.instT' ct pat) ct - fun mk_clist T = apply2 (Thm.cterm_of @{theory}) + fun mk_clist T = apply2 (Thm.global_cterm_of @{theory}) (HOLogic.cons_const T, HOLogic.nil_const T) fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil val mk_pat_list = mk_list (mk_clist @{typ pattern}) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_smt_real.ML --- a/src/HOL/Library/Old_SMT/old_smt_real.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_smt_real.ML Fri Mar 06 15:58:56 2015 +0100 @@ -64,14 +64,14 @@ fun mk_nary _ cu [] = cu | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts) - val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (real)}) - val add = Thm.cterm_of @{theory} @{const plus (real)} + val mk_uminus = Thm.apply (Thm.global_cterm_of @{theory} @{const uminus (real)}) + val add = Thm.global_cterm_of @{theory} @{const plus (real)} val real0 = Numeral.mk_cnumber @{ctyp real} 0 - val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (real)}) - val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (real)}) - val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const divide (real)}) - val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (real)}) - val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (real)}) + val mk_sub = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const minus (real)}) + val mk_mul = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const times (real)}) + val mk_div = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const divide (real)}) + val mk_lt = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less (real)}) + val mk_le = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less_eq (real)}) fun z3_mk_builtin_fun (Old_Z3_Interface.Sym ("-", _)) [ct] = SOME (mk_uminus ct) | z3_mk_builtin_fun (Old_Z3_Interface.Sym ("+", _)) cts = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_smt_solver.ML --- a/src/HOL/Library/Old_SMT/old_smt_solver.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_smt_solver.ML Fri Mar 06 15:58:56 2015 +0100 @@ -195,7 +195,7 @@ const_defs = us}) end) - val cfalse = Thm.cterm_of @{theory} (@{const Trueprop} $ @{const False}) + val cfalse = Thm.global_cterm_of @{theory} (@{const Trueprop} $ @{const False}) in fun add_solver cfg = @@ -268,7 +268,7 @@ (* filter *) -val cnot = Thm.cterm_of @{theory} @{const Not} +val cnot = Thm.global_cterm_of @{theory} @{const Not} fun mk_result outcome xrules = { outcome = outcome, used_facts = xrules } diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_smt_utils.ML --- a/src/HOL/Library/Old_SMT/old_smt_utils.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_smt_utils.ML Fri Mar 06 15:58:56 2015 +0100 @@ -146,7 +146,7 @@ (* patterns and instantiations *) fun mk_const_pat thy name destT = - let val cpat = Thm.cterm_of thy (Const (name, Sign.the_const_type thy name)) + let val cpat = Thm.global_cterm_of thy (Const (name, Sign.the_const_type thy name)) in (destT (Thm.ctyp_of_cterm cpat), cpat) end val destT1 = hd o Thm.dest_ctyp @@ -175,7 +175,7 @@ val dest_all_cbinders = repeat_yield (try o dest_cbinder) -val mk_cprop = Thm.apply (Thm.cterm_of @{theory} @{const Trueprop}) +val mk_cprop = Thm.apply (Thm.global_cterm_of @{theory} @{const Trueprop}) fun dest_cprop ct = (case Thm.term_of ct of diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_interface.ML --- a/src/HOL/Library/Old_SMT/old_z3_interface.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_interface.ML Fri Mar 06 15:58:56 2015 +0100 @@ -137,13 +137,13 @@ | mk_builtin_num ctxt i T = chained_mk_builtin_num ctxt (get_mk_builtins ctxt) i T -val mk_true = Thm.cterm_of @{theory} (@{const Not} $ @{const False}) -val mk_false = Thm.cterm_of @{theory} @{const False} -val mk_not = Thm.apply (Thm.cterm_of @{theory} @{const Not}) -val mk_implies = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.implies}) -val mk_iff = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.eq (bool)}) -val conj = Thm.cterm_of @{theory} @{const HOL.conj} -val disj = Thm.cterm_of @{theory} @{const HOL.disj} +val mk_true = Thm.global_cterm_of @{theory} (@{const Not} $ @{const False}) +val mk_false = Thm.global_cterm_of @{theory} @{const False} +val mk_not = Thm.apply (Thm.global_cterm_of @{theory} @{const Not}) +val mk_implies = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const HOL.implies}) +val mk_iff = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const HOL.eq (bool)}) +val conj = Thm.global_cterm_of @{theory} @{const HOL.conj} +val disj = Thm.global_cterm_of @{theory} @{const HOL.disj} fun mk_nary _ cu [] = cu | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts) @@ -184,15 +184,15 @@ Thm.apply (Thm.mk_binop (Old_SMT_Utils.instTs cTs update) array index) value end -val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (int)}) -val add = Thm.cterm_of @{theory} @{const plus (int)} +val mk_uminus = Thm.apply (Thm.global_cterm_of @{theory} @{const uminus (int)}) +val add = Thm.global_cterm_of @{theory} @{const plus (int)} val int0 = Numeral.mk_cnumber @{ctyp int} 0 -val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (int)}) -val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (int)}) -val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3div}) -val mk_mod = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3mod}) -val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (int)}) -val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (int)}) +val mk_sub = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const minus (int)}) +val mk_mul = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const times (int)}) +val mk_div = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const z3div}) +val mk_mod = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const z3mod}) +val mk_lt = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less (int)}) +val mk_le = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less_eq (int)}) fun mk_builtin_fun ctxt sym cts = (case (sym, cts) of diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_proof_literals.ML --- a/src/HOL/Library/Old_SMT/old_z3_proof_literals.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_proof_literals.ML Fri Mar 06 15:58:56 2015 +0100 @@ -82,7 +82,7 @@ | NONE => false) in exists end -val negate = Thm.apply (Thm.cterm_of @{theory} @{const Not}) +val negate = Thm.apply (Thm.global_cterm_of @{theory} @{const Not}) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_proof_methods.ML --- a/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_proof_methods.ML Fri Mar 06 15:58:56 2015 +0100 @@ -51,15 +51,15 @@ fun mk_inv_of ctxt ct = let val (dT, rT) = Term.dest_funT (Thm.typ_of_cterm ct) - val inv = Proof_Context.cterm_of ctxt (mk_inv_into dT rT) - val univ = Proof_Context.cterm_of ctxt (mk_univ dT) + val inv = Thm.cterm_of ctxt (mk_inv_into dT rT) + val univ = Thm.cterm_of ctxt (mk_univ dT) in Thm.mk_binop inv univ ct end fun mk_inj_prop ctxt ct = let val (dT, rT) = Term.dest_funT (Thm.typ_of_cterm ct) - val inj = Proof_Context.cterm_of ctxt (mk_inj_on dT rT) - val univ = Proof_Context.cterm_of ctxt (mk_univ dT) + val inj = Thm.cterm_of ctxt (mk_inj_on dT rT) + val univ = Thm.cterm_of ctxt (mk_univ dT) in Old_SMT_Utils.mk_cprop (Thm.mk_binop inj ct univ) end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_proof_parser.ML --- a/src/HOL/Library/Old_SMT/old_z3_proof_parser.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_proof_parser.ML Fri Mar 06 15:58:56 2015 +0100 @@ -109,7 +109,7 @@ val max = fold (Integer.max o fst) vars 0 val ns = fst (Variable.variant_fixes (replicate (max + 1) var_prefix) ctxt) fun mk (i, v) = - (v, Proof_Context.cterm_of ctxt (Free (nth ns i, Thm.typ_of_cterm v))) + (v, Thm.cterm_of ctxt (Free (nth ns i, Thm.typ_of_cterm v))) in map mk vars end fun close ctxt (ct, vars) = @@ -120,7 +120,7 @@ fun mk_bound ctxt (i, T) = - let val ct = Proof_Context.cterm_of ctxt (Var ((Name.uu, 0), T)) + let val ct = Thm.cterm_of ctxt (Var ((Name.uu, 0), T)) in (ct, [(i, ct)]) end local @@ -129,7 +129,7 @@ val cv = (case AList.lookup (op =) vars 0 of SOME cv => cv - | _ => Proof_Context.cterm_of ctxt (Var ((Name.uu, Thm.maxidx_of_cterm ct + 1), T))) + | _ => Thm.cterm_of ctxt (Var ((Name.uu, Thm.maxidx_of_cterm ct + 1), T))) fun dec (i, v) = if i = 0 then NONE else SOME (i-1, v) val vars' = map_filter dec vars in (Thm.apply (Old_SMT_Utils.instT' cv q) (Thm.lambda cv ct), vars') end @@ -181,7 +181,7 @@ (** parser context **) -val not_false = Thm.cterm_of @{theory} (@{const Not} $ @{const False}) +val not_false = Thm.global_cterm_of @{theory} (@{const Not} $ @{const False}) fun make_context ctxt typs terms = let @@ -191,7 +191,7 @@ |> Symtab.fold (Variable.declare_term o snd) terms fun cert @{const True} = not_false - | cert t = Proof_Context.cterm_of ctxt' t + | cert t = Thm.cterm_of ctxt' t in (typs, Symtab.map (K cert) terms, Inttab.empty, [], [], ctxt') end @@ -207,7 +207,7 @@ | NONE => cx |> fresh_name (decl_prefix ^ n) |> (fn (m, (typs, terms, exprs, steps, vars, ctxt)) => let - val upd = Symtab.update (n, Proof_Context.cterm_of ctxt (Free (m, T))) + val upd = Symtab.update (n, Thm.cterm_of ctxt (Free (m, T))) in (typs, upd terms, exprs, steps, vars, ctxt) end)) fun mk_typ (typs, _, _, _, _, ctxt) (s as Old_Z3_Interface.Sym (n, _)) = @@ -381,7 +381,7 @@ fun pats st = seps (par ((this ":pat" || this ":nopat") |-- seps1 id)) st -val ctrue = Thm.cterm_of @{theory} @{const True} +val ctrue = Thm.global_cterm_of @{theory} @{const True} fun pattern st = par (this "pattern" |-- Scan.repeat1 (sep arg) >> (the o mk_fun (K (SOME ctrue)))) st diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML --- a/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML Fri Mar 06 15:58:56 2015 +0100 @@ -617,7 +617,7 @@ fun get_vars f mk pred ctxt t = Term.fold_aterms f t [] |> map_filter (fn v => - if pred v then SOME (Proof_Context.cterm_of ctxt (mk v)) else NONE) + if pred v then SOME (Thm.cterm_of ctxt (mk v)) else NONE) fun close vars f ct ctxt = let diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Old_SMT/old_z3_proof_tools.ML --- a/src/HOL/Library/Old_SMT/old_z3_proof_tools.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Old_SMT/old_z3_proof_tools.ML Fri Mar 06 15:58:56 2015 +0100 @@ -156,7 +156,7 @@ | NONE => let val (n, ctxt') = yield_singleton Variable.variant_fixes "x" ctxt - val cv = Proof_Context.cterm_of ctxt' + val cv = Thm.cterm_of ctxt' (Free (n, map Thm.typ_of_cterm cvs' ---> Thm.typ_of_cterm ct)) val cu = Drule.list_comb (cv, cvs') val e = (t, (cv, fold_rev Thm.lambda cvs' ct)) @@ -229,7 +229,7 @@ | _ => fresh_abstraction dcvs ct cx))) in abstr (depth, []) end -val cimp = Thm.cterm_of @{theory} @{const Pure.imp} +val cimp = Thm.global_cterm_of @{theory} @{const Pure.imp} fun deepen depth f x = if depth = 0 then f depth x diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML --- a/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz_tools.ML Fri Mar 06 15:58:56 2015 +0100 @@ -114,7 +114,7 @@ val parse_id = Scan.one Symbol.is_letter ::: Scan.many Symbol.is_letdig >> implode fun parse_varpow ctxt = parse_id -- Scan.optional (str "^" |-- nat) 1 >> - (fn (x, k) => (Proof_Context.cterm_of ctxt (Free (x, @{typ real})), k)) + (fn (x, k) => (Thm.cterm_of ctxt (Free (x, @{typ real})), k)) fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >> (fn xs => fold FuncUtil.Ctermfunc.update xs FuncUtil.Ctermfunc.empty) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Library/simps_case_conv.ML --- a/src/HOL/Library/simps_case_conv.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Library/simps_case_conv.ML Fri Mar 06 15:58:56 2015 +0100 @@ -97,7 +97,7 @@ | import (thm :: thms) ctxt = let val fun_ct = strip_eq #> fst #> strip_comb #> fst #> Logic.mk_term - #> Proof_Context.cterm_of ctxt + #> Thm.cterm_of ctxt val ct = fun_ct thm val cts = map fun_ct thms val pairs = map (fn s => (s,ct)) cts diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Matrix_LP/Compute_Oracle/compute.ML --- a/src/HOL/Matrix_LP/Compute_Oracle/compute.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML Fri Mar 06 15:58:56 2015 +0100 @@ -362,13 +362,13 @@ commas (map (Syntax.string_of_sort_global thy) shyps)) else () in - Thm.cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop) + Thm.global_cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop) end))); fun export_thm thy hyps shyps prop = let val th = export_oracle (thy, hyps, shyps, prop) - val hyps = map (fn h => Thm.assume (Thm.cterm_of thy h)) hyps + val hyps = map (fn h => Thm.assume (Thm.global_cterm_of thy h)) hyps in fold (fn h => fn p => Thm.implies_elim p h) hyps th end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Matrix_LP/Compute_Oracle/linker.ML --- a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML Fri Mar 06 15:58:56 2015 +0100 @@ -315,7 +315,7 @@ end fun conv_subst thy (subst : Type.tyenv) = - map (fn (iname, (sort, ty)) => (Thm.ctyp_of thy (TVar (iname, sort)), Thm.ctyp_of thy ty)) + map (fn (iname, (sort, ty)) => (Thm.global_ctyp_of thy (TVar (iname, sort)), Thm.global_ctyp_of thy ty)) (Vartab.dest subst) fun add_monos thy monocs pats ths = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Multivariate_Analysis/normarith.ML --- a/src/HOL/Multivariate_Analysis/normarith.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/normarith.ML Fri Mar 06 15:58:56 2015 +0100 @@ -344,7 +344,7 @@ fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms) fun mk_norm t = Thm.apply (instantiate_cterm' [SOME (Thm.ctyp_of_cterm t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t fun mk_equals l r = Thm.apply (Thm.apply (instantiate_cterm' [SOME (Thm.ctyp_of_cterm 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 asl = map2 (fn (t,_) => fn n => Thm.assume (mk_equals (mk_norm t) (Thm.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)) val ges' = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Mutabelle/mutabelle_extra.ML --- a/src/HOL/Mutabelle/mutabelle_extra.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Mutabelle/mutabelle_extra.ML Fri Mar 06 15:58:56 2015 +0100 @@ -401,7 +401,7 @@ |> map Mutabelle.freeze |> map freezeT (* |> filter (not o is_forbidden_mutant) *) |> map_filter (try (Sign.cert_term thy)) - |> filter (is_some o try (Thm.cterm_of thy)) + |> filter (is_some o try (Thm.global_cterm_of thy)) |> filter (is_some o try (Syntax.check_term (Proof_Context.init_global thy))) |> take_random max_mutants val _ = map (fn t => writeln ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/NSA/transfer.ML --- a/src/HOL/NSA/transfer.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/NSA/transfer.ML Fri Mar 06 15:58:56 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') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of thy t)) + val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.global_cterm_of thy t)) val u = unstar_term consts t' val tac = rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_datatype.ML Fri Mar 06 15:58:56 2015 +0100 @@ -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 Thm.cterm_of thy + val cert = SOME o Thm.global_cterm_of thy in - SOME (mk_meta_eq (Drule.instantiate' [SOME (Thm.ctyp_of thy S)] + SOME (mk_meta_eq (Drule.instantiate' [SOME (Thm.global_ctyp_of thy S)] [cert t, cert r, cert s] dj_cp')) end else NONE @@ -777,12 +777,12 @@ fun dt_constr_defs ((((((_, (_, _, constrs)), (_, (_, _, constrs'))), tname), T), T'), constr_syntax) (thy, defs, eqns, dist_lemmas) = let - val rep_const = Thm.cterm_of thy + val rep_const = Thm.global_cterm_of thy (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); val dist = Drule.export_without_context (cterm_instantiate - [(Thm.cterm_of thy distinct_f, rep_const)] Old_Datatype.distinct_lemma); + [(Thm.global_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 @@ -1051,7 +1051,7 @@ 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 (Thm.cterm_of thy8) Ps ~~ map (Thm.cterm_of thy8) frees) indrule_lemma; + (map (Thm.global_cterm_of thy8) Ps ~~ map (Thm.global_cterm_of thy8) frees) indrule_lemma; val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; @@ -1268,7 +1268,7 @@ | _ => false)) perm_fresh_fresh in Drule.instantiate' [] - [SOME (Thm.cterm_of thy a), NONE, SOME (Thm.cterm_of thy b)] th + [SOME (Thm.global_cterm_of thy a), NONE, SOME (Thm.global_cterm_of thy b)] th end; val fs_cp_sort = @@ -1382,13 +1382,13 @@ val induct_aux' = Thm.instantiate ([], map (fn (s, v as Var (_, T)) => - (Thm.cterm_of thy9 v, Thm.cterm_of thy9 (Free (s, T)))) + (Thm.global_cterm_of thy9 v, Thm.global_cterm_of thy9 (Free (s, T)))) (pnames ~~ map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct_aux)))) @ map (fn (_, f) => let val f' = Logic.varify_global f - in (Thm.cterm_of thy9 f', - Thm.cterm_of thy9 (Const (@{const_name Nominal.supp}, fastype_of f'))) + in (Thm.global_cterm_of thy9 f', + Thm.global_cterm_of thy9 (Const (@{const_name Nominal.supp}, fastype_of f'))) end) fresh_fs) induct_aux; val induct = Goal.prove_global_future thy9 [] @@ -1551,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 ([], - [(Thm.cterm_of thy11 (Var (("pi", 0), permT)), - Thm.cterm_of thy11 (Const (@{const_name rev}, permT --> permT) $ pi))]) th) 1 THEN + [(Thm.global_cterm_of thy11 (Var (("pi", 0), permT)), + Thm.global_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); @@ -1630,8 +1630,8 @@ val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns') in EVERY [rtac (Drule.cterm_instantiate - [(Thm.cterm_of thy11 S, - Thm.cterm_of thy11 (Const (@{const_name Nominal.supp}, + [(Thm.global_cterm_of thy11 S, + Thm.global_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 @@ -1642,8 +1642,8 @@ SUBPROOF (fn {prems = prems', params = [(_, a), (_, b)], ...} => EVERY [cut_facts_tac [rec_prem] 1, rtac (Thm.instantiate ([], - [(Thm.cterm_of thy11 (Var (("pi", 0), mk_permT aT)), - Thm.cterm_of thy11 + [(Thm.global_cterm_of thy11 (Var (("pi", 0), mk_permT aT)), + Thm.global_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, @@ -1671,7 +1671,7 @@ (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets); val induct_aux_rec = Drule.cterm_instantiate - (map (apply2 (Thm.cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort)) + (map (apply2 (Thm.global_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 @ @@ -1861,8 +1861,8 @@ val l = find_index (equal T) dt_atomTs; val th = nth (nth rec_equiv_thms' l) k; val th' = Thm.instantiate ([], - [(Thm.cterm_of thy11 (Var (("pi", 0), U)), - Thm.cterm_of thy11 p)]) th; + [(Thm.global_cterm_of thy11 (Var (("pi", 0), U)), + Thm.global_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)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_fresh_fun.ML --- a/src/HOL/Nominal/nominal_fresh_fun.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_fresh_fun.ML Fri Mar 06 15:58:56 2015 +0100 @@ -23,8 +23,8 @@ (head_of (Logic.incr_indexes (Ts, j) t), fold_rev Term.abs ps u)) tinst; val th' = instf - (map (apply2 (Thm.ctyp_of thy)) tyinst') - (map (apply2 (Thm.cterm_of thy)) tinst') + (map (apply2 (Thm.global_ctyp_of thy)) tyinst') + (map (apply2 (Thm.global_cterm_of thy)) tinst') (Thm.lift_rule cgoal th) in compose_tac ctxt (elim, th', Thm.nprems_of th) i st @@ -149,7 +149,7 @@ val thy = Thm.theory_of_thm st; in case subtract (op =) vars vars' of [x] => - Seq.single (Thm.instantiate ([], [(Thm.cterm_of thy x, Thm.cterm_of thy (fold_rev Term.abs params (Bound 0)))]) st) + Seq.single (Thm.instantiate ([], [(Thm.global_cterm_of thy x, Thm.global_cterm_of thy (fold_rev Term.abs params (Bound 0)))]) st) | _ => error "fresh_fun_simp: Too many variables, please report." end in diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_induct.ML --- a/src/HOL/Nominal/nominal_induct.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_induct.ML Fri Mar 06 15:58:56 2015 +0100 @@ -54,7 +54,7 @@ end; val substs = map2 subst insts concls |> flat |> distinct (op =) - |> map (apply2 (Proof_Context.cterm_of ctxt)); + |> map (apply2 (Thm.cterm_of ctxt)); in (((cases, nconcls), consumes), Drule.cterm_instantiate substs joined_rule) end; @@ -84,7 +84,7 @@ fun nominal_induct_tac ctxt simp def_insts avoiding fixings rules facts = let val thy = Proof_Context.theory_of ctxt; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val ((insts, defs), defs_ctxt) = fold_map Induct.add_defs def_insts ctxt |>> split_list; val atomized_defs = map (map (Conv.fconv_rule (Induct.atomize_cterm ctxt))) defs; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_inductive.ML Fri Mar 06 15:58:56 2015 +0100 @@ -204,7 +204,7 @@ val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Drule.instantiate' - [SOME (Thm.ctyp_of thy fsT)] [] inductive_forall_def; + [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); @@ -338,7 +338,7 @@ val pis'' = fold (concat_perm #> map) pis' pis; val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) (Vartab.empty, Vartab.empty); - val ihyp' = Thm.instantiate ([], map (apply2 (Thm.cterm_of thy)) + val ihyp' = Thm.instantiate ([], map (apply2 (Thm.global_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 Thm.cterm_of thy) + (fold_rev (mk_perm_bool o Thm.global_cterm_of thy) (rev pis' @ pis) th)); val (gprems1, gprems2) = split_list (map (fn (th, t) => @@ -505,8 +505,8 @@ 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 (Thm.cterm_of thy) pis); - val obj = Thm.cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms + val mk_pis = fold_rev mk_perm_bool (map (Thm.global_cterm_of thy) pis); + val obj = Thm.global_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 @@ -630,9 +630,9 @@ 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 (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) + (mk_perm_bool (Thm.global_cterm_of thy pi) th)) prems'; + val intr' = Drule.cterm_instantiate (map (Thm.global_cterm_of thy) vs ~~ + map (Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Fri Mar 06 15:58:56 2015 +0100 @@ -147,7 +147,7 @@ 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 #> Thm.cterm_of thy) vs ~~ cts) th + map (Envir.subst_term env #> Thm.global_cterm_of thy) vs ~~ cts) th end; fun prove_strong_ind s alt_name avoids ctxt = @@ -230,7 +230,7 @@ val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Drule.instantiate' - [SOME (Thm.ctyp_of thy fsT)] [] inductive_forall_def; + [SOME (Thm.global_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 (Thm.ctyp_of thy (fastype_of p))] [SOME (Thm.cterm_of thy p)] + [SOME (Thm.global_ctyp_of thy (fastype_of p))] [SOME (Thm.global_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 @@ -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) #> Thm.cterm_of thy) params' @ [Thm.cterm_of thy z]); + (pis' @ pis) #> Thm.global_cterm_of thy) params' @ [Thm.global_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 Thm.cterm_of thy) + (fold_rev (mk_perm_bool o Thm.global_cterm_of thy) (pis' @ pis) th)); val gprems2 = map (fn (th, t) => if null (preds_of ps t) then mk_pi th diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_permeq.ML --- a/src/HOL/Nominal/nominal_permeq.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_permeq.ML Fri Mar 06 15:58:56 2015 +0100 @@ -203,14 +203,14 @@ if pi1 <> pi2 then (* only apply the composition rule in this case *) if T = U then SOME (Drule.instantiate' - [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)] + [SOME (Thm.global_ctyp_of thy (fastype_of t))] + [SOME (Thm.global_cterm_of thy pi1), SOME (Thm.global_cterm_of thy pi2), SOME (Thm.global_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 (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)] + [SOME (Thm.global_ctyp_of thy (fastype_of t))] + [SOME (Thm.global_cterm_of thy pi1), SOME (Thm.global_cterm_of thy pi2), SOME (Thm.global_cterm_of thy t)] (mk_meta_eq (Global_Theory.get_thm thy ("cp_"^tname'^"_"^uname'^"_inst") RS cp1_aux))) else NONE @@ -296,7 +296,7 @@ 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 cert = Thm.global_cterm_of (Thm.theory_of_thm st); val ps = Logic.strip_params (Thm.term_of goal); val Ts = rev (map snd ps); val vs = collect_vars 0 x []; @@ -340,7 +340,7 @@ 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 cert = Thm.global_cterm_of (Thm.theory_of_thm st); val ps = Logic.strip_params (Thm.term_of goal); val Ts = rev (map snd ps); val vs = collect_vars 0 t []; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Fri Mar 06 15:58:56 2015 +0100 @@ -281,7 +281,7 @@ val qualify = Binding.qualify false (space_implode "_" (map (Long_Name.base_name o #1) defs)); val names_atts' = map (apfst qualify) names_atts; - val cert = Proof_Context.cterm_of lthy'; + val cert = Thm.cterm_of lthy'; fun mk_idx eq = let diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Nominal/nominal_thmdecls.ML --- a/src/HOL/Nominal/nominal_thmdecls.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Nominal/nominal_thmdecls.ML Fri Mar 06 15:58:56 2015 +0100 @@ -53,9 +53,9 @@ fun prove_eqvt_tac ctxt orig_thm pi pi' = let val thy = Proof_Context.theory_of ctxt - val mypi = Thm.cterm_of thy pi + val mypi = Thm.global_cterm_of thy pi val T = fastype_of pi' - val mypifree = Thm.cterm_of thy (Const (@{const_name "rev"}, T --> T) $ pi') + val mypifree = Thm.global_cterm_of thy (Const (@{const_name "rev"}, T --> T) $ pi') val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp" in EVERY1 [tactic ctxt ("iffI applied", rtac @{thm iffI}), diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Num.thy --- a/src/HOL/Num.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Num.thy Fri Mar 06 15:58:56 2015 +0100 @@ -1198,7 +1198,7 @@ fun number_of thy T n = if not (Sign.of_sort thy (T, @{sort numeral})) then raise CTERM ("number_of", []) - else Numeral.mk_cnumber (Thm.ctyp_of thy T) n; + else Numeral.mk_cnumber (Thm.global_ctyp_of thy T) n; in K ( Lin_Arith.add_simps (@{thms arith_simps} @ @{thms more_arith_simps} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Probability/measurable.ML --- a/src/HOL/Probability/measurable.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Probability/measurable.ML Fri Mar 06 15:58:56 2015 +0100 @@ -251,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 Thm.ctyp_of Ts) (cert Thm.cterm_of ts) + Drule.instantiate' (cert Thm.global_ctyp_of Ts) (cert Thm.global_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) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Proofs/Lambda/WeakNorm.thy --- a/src/HOL/Proofs/Lambda/WeakNorm.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Proofs/Lambda/WeakNorm.thy Fri Mar 06 15:58:56 2015 +0100 @@ -433,11 +433,11 @@ 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 (Thm.term_of ct1)); -val ct1' = Thm.cterm_of @{theory} (term_of_dB [] (Thm.typ_of_cterm ct1) dB1); +val ct1' = Thm.global_cterm_of @{theory} (term_of_dB [] (Thm.typ_of_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 (Thm.term_of ct2)); -val ct2' = Thm.cterm_of @{theory} (term_of_dB [] (Thm.typ_of_cterm ct2) dB2); +val ct2' = Thm.global_cterm_of @{theory} (term_of_dB [] (Thm.typ_of_cterm ct2) dB2); *} end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Statespace/distinct_tree_prover.ML --- a/src/HOL/Statespace/distinct_tree_prover.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Statespace/distinct_tree_prover.ML Fri Mar 06 15:58:56 2015 +0100 @@ -95,7 +95,7 @@ fun mapT_and_recertify ct = let 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; + in (Thm.global_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; @@ -150,10 +150,10 @@ val tyinsts' = map (fn (v, (S, U)) => - (Thm.ctyp_of thy (TVar (v, S)), Thm.ctyp_of thy U)) tyinsts; + (Thm.global_ctyp_of thy (TVar (v, S)), Thm.global_ctyp_of thy U)) tyinsts; val insts' = map (fn (idxn, ct) => - (Thm.cterm_of thy (Var (idxn, Thm.typ_of_cterm ct)), ct)) insts; + (Thm.global_cterm_of thy (Var (idxn, Thm.typ_of_cterm ct)), ct)) insts; val rule' = Thm.instantiate (tyinsts', insts') rule; in fold Thm.elim_implies prems rule' end; @@ -298,8 +298,8 @@ val [alphaI] = #2 (dest_Type T); in Thm.instantiate - ([(alpha, Thm.ctyp_of thy alphaI)], - [(Thm.cterm_of thy (Var (v, treeT alphaI)), ct')]) @{thm subtract_Tip} + ([(alpha, Thm.global_ctyp_of thy alphaI)], + [(Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Statespace/state_fun.ML --- a/src/HOL/Statespace/state_fun.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Statespace/state_fun.ML Fri Mar 06 15:58:56 2015 +0100 @@ -56,13 +56,13 @@ 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 (Thm.cterm_of thy P); + val P_P' = Simplifier.rewrite ctxt (Thm.global_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 (Thm.cterm_of thy Q); + val Q_Q' = Simplifier.rewrite ctxt (Thm.global_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']) @@ -141,7 +141,7 @@ | mk_upds s = (Var (("s", mi + 1), sT), mi + 2); val ct = - Thm.cterm_of thy (Const (@{const_name StateFun.lookup}, lT) $ destr $ n $ fst (mk_upds s)); + Thm.global_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; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/String.thy --- a/src/HOL/String.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/String.thy Fri Mar 06 15:58:56 2015 +0100 @@ -248,7 +248,7 @@ setup {* let - val nibbles = map_range (Thm.cterm_of @{theory} o HOLogic.mk_nibble) 16; + val nibbles = map_range (Thm.global_cterm_of @{theory} o HOLogic.mk_nibble) 16; val simpset = put_simpset HOL_ss @{context} addsimps @{thms nat_of_nibble.simps mult_0 mult_1 add_0 add_0_right arith_simps numeral_plus_one}; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML --- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Fri Mar 06 15:58:56 2015 +0100 @@ -572,14 +572,14 @@ diff thy (scheme_t, instance_t) (*valuation of type variables*) - val typeval = map (apply2 (Thm.ctyp_of thy)) type_pairing + val typeval = map (apply2 (Thm.global_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 (Thm.cterm_of thy)) + |> map (apply2 (Thm.global_cterm_of thy)) in Thm.instantiate (typeval, termval) scheme_thm end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/TPTP/TPTP_Proof_Reconstruction.thy --- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Fri Mar 06 15:58:56 2015 +0100 @@ -917,7 +917,7 @@ (* val contextualise = fold absdummy (map snd params) *) val contextualise = fold absfree params - val skolem_cts = map (contextualise #> Thm.cterm_of thy) skolem_terms + val skolem_cts = map (contextualise #> Thm.global_cterm_of thy) skolem_terms (*now the instantiation code*) @@ -937,7 +937,7 @@ fun make_var pre_var = the_single pre_var |> Var - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy |> SOME in if null pre_var then NONE @@ -2003,7 +2003,7 @@ thy prob_name (#meta pannot) n |> the - |> (fn {inference_fmla, ...} => Thm.cterm_of thy inference_fmla) + |> (fn {inference_fmla, ...} => Thm.global_cterm_of thy inference_fmla) |> oracle_iinterp end *} diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/ATP/atp_problem_generate.ML --- a/src/HOL/Tools/ATP/atp_problem_generate.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/ATP/atp_problem_generate.ML Fri Mar 06 15:58:56 2015 +0100 @@ -760,7 +760,7 @@ fun do_introduce_combinators ctxt Ts t = let val thy = Proof_Context.theory_of ctxt in t |> conceal_bounds Ts - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy |> Meson_Clausify.introduce_combinators_in_cterm |> Thm.prop_of |> Logic.dest_equals |> snd |> reveal_bounds Ts diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/ATP/atp_util.ML --- a/src/HOL/Tools/ATP/atp_util.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/ATP/atp_util.ML Fri Mar 06 15:58:56 2015 +0100 @@ -353,7 +353,7 @@ fun abs_extensionalize_term ctxt t = if exists_Const is_fun_equality t then let val thy = Proof_Context.theory_of ctxt in - t |> Thm.cterm_of thy |> Meson.abs_extensionalize_conv ctxt + t |> Thm.global_cterm_of thy |> Meson.abs_extensionalize_conv ctxt |> Thm.prop_of |> Logic.dest_equals |> snd end else diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_def.ML Fri Mar 06 15:58:56 2015 +0100 @@ -1227,11 +1227,11 @@ val funTs = map (fn T => bdT --> T) ranTs; val ran_bnfT = mk_bnf_T ranTs Calpha; val (revTs, Ts) = `rev (bd_bnfT :: funTs); - val cTs = map (SOME o Proof_Context.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; + val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs) (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, map Bound (live - 1 downto 0)) $ Bound live)); - val cts = [NONE, SOME (Proof_Context.cterm_of lthy tinst)]; + val cts = [NONE, SOME (Thm.cterm_of lthy tinst)]; in Drule.instantiate' cTs cts @{thm surj_imp_ordLeq} end); @@ -1346,7 +1346,7 @@ fun mk_rel_flip () = let val rel_conversep_thm = Lazy.force rel_conversep; - val cts = map (SOME o Proof_Context.cterm_of lthy) Rs; + val cts = map (SOME o Thm.cterm_of lthy) Rs; val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm; in unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD}) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Fri Mar 06 15:58:56 2015 +0100 @@ -476,7 +476,7 @@ let 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'; + val cRs = map (fn f => (Thm.cterm_of lthy (Var f), Thm.cterm_of lthy (mk_flip f))) Rs'; in Drule.cterm_instantiate cRs thm end; @@ -598,14 +598,14 @@ Drule.dummy_thm else let val ctor' = mk_ctor Bs ctor in - cterm_instantiate_pos [NONE, NONE, SOME (Proof_Context.cterm_of lthy ctor')] arg_cong + cterm_instantiate_pos [NONE, NONE, SOME (Thm.cterm_of lthy ctor')] arg_cong end; fun mk_cIn ctor k xs = let val absT = domain_type (fastype_of ctor) in mk_absumprod absT abs n k xs |> fp = Greatest_FP ? curry (op $) ctor - |> Proof_Context.cterm_of lthy + |> Thm.cterm_of lthy end; val cxIns = map2 (mk_cIn ctor) ks xss; @@ -615,7 +615,7 @@ fold_thms lthy [ctr_def'] (unfold_thms lthy (o_apply :: pre_map_def :: (if fp = Least_FP then [] else [dtor_ctor]) @ sumprod_thms_map @ abs_inverses) - (cterm_instantiate_pos (map (SOME o Proof_Context.cterm_of lthy) fs @ [SOME cxIn]) + (cterm_instantiate_pos (map (SOME o Thm.cterm_of lthy) fs @ [SOME cxIn]) (if fp = Least_FP then fp_map_thm else fp_map_thm RS ctor_cong RS (ctor_dtor RS sym RS trans)))) |> singleton (Proof_Context.export names_lthy lthy); @@ -643,7 +643,7 @@ (unfold_thms lthy (pre_rel_def :: abs_inverses @ (if fp = Least_FP then [] else [dtor_ctor]) @ sumprod_thms_rel @ @{thms vimage2p_def sum.inject sum.distinct(1)[THEN eq_False[THEN iffD2]]}) - (cterm_instantiate_pos (map (SOME o Proof_Context.cterm_of lthy) Rs @ [SOME cxIn, SOME cyIn]) + (cterm_instantiate_pos (map (SOME o Thm.cterm_of lthy) Rs @ [SOME cxIn, SOME cyIn]) fp_rel_thm)) |> postproc |> singleton (Proof_Context.export names_lthy lthy); @@ -734,7 +734,7 @@ val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); val thm = Goal.prove_sorry lthy [] (flat premss) goal (fn {context = ctxt, prems} => - mk_set_cases_tac ctxt (Proof_Context.cterm_of ctxt ta) prems exhaust set_thms) + mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation |> rotate_prems ~1; @@ -812,7 +812,7 @@ fun prove goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => - mk_rel_sel_tac ctxt (Proof_Context.cterm_of ctxt ta) (Proof_Context.cterm_of ctxt tb) exhaust (flat disc_thmss) + mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation; @@ -846,7 +846,7 @@ val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); val thm = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => - mk_rel_cases_tac ctxt (Proof_Context.cterm_of ctxt ta) (Proof_Context.cterm_of ctxt tb) exhaust injects + mk_rel_cases_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation; @@ -920,7 +920,7 @@ else Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) (fn {context = ctxt, prems = _} => - mk_map_disc_iff_tac ctxt (Proof_Context.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) + mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) |> Conjunction.elim_balanced (length goals) |> Proof_Context.export names_lthy lthy |> map Thm.close_derivation @@ -954,7 +954,7 @@ else Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) (fn {context = ctxt, prems = _} => - mk_map_sel_tac ctxt (Proof_Context.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms + mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms (flat sel_thmss) live_nesting_map_id0s) |> Conjunction.elim_balanced (length goals) |> Proof_Context.export names_lthy lthy @@ -1013,7 +1013,7 @@ else Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) (fn {context = ctxt, prems = _} => - mk_set_sel_tac ctxt (Proof_Context.cterm_of ctxt ta) exhaust (flat disc_thmss) + mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) set0_thms) |> Conjunction.elim_balanced (length goals) |> Proof_Context.export names_lthy lthy @@ -1302,7 +1302,7 @@ val rel_induct0_thm = Goal.prove_sorry lthy [] prems goal (fn {context = ctxt, prems} => - mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Proof_Context.cterm_of ctxt) IRs) exhausts ctor_defss + mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation; @@ -1560,7 +1560,7 @@ val rel_coinduct0_thm = Goal.prove_sorry lthy [] prems goal (fn {context = ctxt, prems} => - mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Proof_Context.cterm_of ctxt) IRs) prems + mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars) (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses live_nesting_rel_eqs) @@ -1644,7 +1644,7 @@ val thm = Goal.prove_sorry lthy [] prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => - mk_set_induct0_tac ctxt (map (Proof_Context.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts + mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> singleton (Proof_Context.export ctxt'' ctxt) |> Thm.close_derivation; @@ -1819,7 +1819,7 @@ val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; - fun mk_case_split' cp = Drule.instantiate' [] [SOME (Proof_Context.cterm_of lthy cp)] @{thm case_split}; + fun mk_case_split' cp = Drule.instantiate' [] [SOME (Thm.cterm_of lthy cp)] @{thm case_split}; val case_splitss' = map (map mk_case_split') cpss; @@ -1845,8 +1845,8 @@ let val (domT, ranT) = dest_funT (fastype_of sel); val arg_cong' = - Drule.instantiate' (map (SOME o Proof_Context.ctyp_of lthy) [domT, ranT]) - [NONE, NONE, SOME (Proof_Context.cterm_of lthy sel)] arg_cong + Drule.instantiate' (map (SOME o Thm.ctyp_of lthy) [domT, ranT]) + [NONE, NONE, SOME (Thm.cterm_of lthy sel)] arg_cong |> Thm.varifyT_global; val sel_thm' = sel_thm RSN (2, trans); in @@ -2141,8 +2141,8 @@ (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => - mk_ctor_iff_dtor_tac ctxt (map (SOME o Proof_Context.ctyp_of lthy) [ctr_absT, fpT]) - (Proof_Context.cterm_of lthy ctor) (Proof_Context.cterm_of lthy dtor) ctor_dtor dtor_ctor) + mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctr_absT, fpT]) + (Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor) |> Morphism.thm phi |> Thm.close_derivation end; @@ -2233,7 +2233,7 @@ let val (Rs, Ss, goals, names_lthy) = mk_co_rec_transfer_goals lthy recs in Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) (fn {context = ctxt, prems = _} => - mk_rec_transfer_tac ctxt nn ns (map (Proof_Context.cterm_of ctxt) Ss) (map (Proof_Context.cterm_of ctxt) Rs) xsssss + mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) |> Conjunction.elim_balanced nn |> Proof_Context.export names_lthy lthy @@ -2385,7 +2385,7 @@ in Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) (fn {context = ctxt, prems = _} => - mk_corec_transfer_tac ctxt (map (Proof_Context.cterm_of ctxt) Ss) (map (Proof_Context.cterm_of ctxt) Rs) + mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs (flat pgss) pss qssss gssss) |> Conjunction.elim_balanced (length goals) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Fri Mar 06 15:58:56 2015 +0100 @@ -97,7 +97,7 @@ 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)); + (Thm.cterm_of ctxt (Var f), Thm.cterm_of ctxt (mk_proj T (num_binder_types T) k)); val cfps = map mk_cfp fs; in Drule.cterm_instantiate cfps thm diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_fp_n2m.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML Fri Mar 06 15:58:56 2015 +0100 @@ -191,7 +191,7 @@ case_fp fp (snd o Term.dest_comb) (snd o Term.dest_comb o fst o Term.dest_comb); val raw_phis = map (extract o HOLogic.dest_Trueprop o Thm.concl_of) rel_xtor_co_inducts; val thetas = AList.group (op =) - (mutual_cliques ~~ map (apply2 (Proof_Context.cterm_of lthy)) (raw_phis ~~ pre_phis)); + (mutual_cliques ~~ map (apply2 (Thm.cterm_of lthy)) (raw_phis ~~ pre_phis)); in map2 (Drule.cterm_instantiate o the o AList.lookup (op =) thetas) mutual_cliques rel_xtor_co_inducts @@ -214,7 +214,7 @@ fun mk_Grp_id P = let val T = domain_type (fastype_of P); in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end; - val cts = map (SOME o Proof_Context.cterm_of lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps); + val cts = map (SOME o Thm.cterm_of lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps); fun mk_fp_type_copy_thms thm = map (curry op RS thm) @{thms type_copy_Abs_o_Rep type_copy_vimage2p_Grp_Rep}; fun mk_type_copy_thms thm = map (curry op RS thm) @@ -235,7 +235,7 @@ end | Greatest_FP => let - val cts = NONE :: map (SOME o Proof_Context.cterm_of lthy) (map HOLogic.eq_const As); + val cts = NONE :: map (SOME o Thm.cterm_of lthy) (map HOLogic.eq_const As); in cterm_instantiate_pos cts xtor_rel_co_induct |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_fp_rec_sugar_transfer.ML --- a/src/HOL/Tools/BNF/bnf_fp_rec_sugar_transfer.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_rec_sugar_transfer.ML Fri Mar 06 15:58:56 2015 +0100 @@ -44,7 +44,7 @@ val else_branch = Var (("e", j), T); val lam = Term.lambda cond (mk_If cond then_branch else_branch); in - cterm_instantiate_pos [SOME (Proof_Context.cterm_of ctxt lam)] thm + cterm_instantiate_pos [SOME (Thm.cterm_of ctxt lam)] thm end; val transfer_rules = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_fp_util.ML --- a/src/HOL/Tools/BNF/bnf_fp_util.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_fp_util.ML Fri Mar 06 15:58:56 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 (Thm.ctyp_of @{theory} T)] [] + |> Drule.instantiate' [SOME (Thm.global_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 @@ -578,7 +578,7 @@ val n = Thm.nprems_of coind; val m = Thm.nprems_of (hd rel_monos) - n; fun mk_inst phi = (phi, mk_union (phi, HOLogic.eq_const (fst (dest_pred2T (fastype_of phi))))) - |> apply2 (Proof_Context.cterm_of ctxt); + |> apply2 (Thm.cterm_of ctxt); val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map (mk_inst o Var); fun mk_unfold rel_eq rel_mono = let diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_gfp.ML --- a/src/HOL/Tools/BNF/bnf_gfp.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_gfp.ML Fri Mar 06 15:58:56 2015 +0100 @@ -1080,7 +1080,7 @@ val goal = list_all_free (kl :: zs) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); - val cts = map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' goal, nat]; + val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val length_Lev = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) @@ -1115,8 +1115,8 @@ Library.foldr1 HOLogic.mk_conj (map2 (mk_conjunct i z) ks zs_copy)) ks zs)); - val cTs = [SOME (Proof_Context.ctyp_of lthy sum_sbdT)]; - val cts = map (SOME o Proof_Context.cterm_of lthy) [Term.absfree kl' goal, kl]; + val cTs = [SOME (Thm.ctyp_of lthy sum_sbdT)]; + val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree kl' goal, kl]; val rv_last = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) @@ -1153,7 +1153,7 @@ val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); - val cts = map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' goal, nat]; + val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_Lev = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) @@ -1192,7 +1192,7 @@ val goal = list_all_free (kl :: k :: zs @ zs_copy) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); - val cts = map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' goal, nat]; + val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_image_Lev = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal) @@ -1867,7 +1867,7 @@ Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls val ctss = - map (fn phi => map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' phi, nat]) concls; + map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) concls; in @{map 4} (fn goal => fn cts => fn col_0s => fn col_Sucs => Goal.prove_sorry lthy [] [] goal @@ -1952,10 +1952,10 @@ maps (map (fn thm => thm RS @{thm subset_Collect_iff})) dtor_Jset_incl_thmss @ @{thms subset_Collect_iff[OF subset_refl]}; - val cTs = map (SOME o Proof_Context.ctyp_of lthy) params'; + val cTs = map (SOME o Thm.ctyp_of lthy) params'; fun mk_induct_tinst phis jsets y y' = @{map 4} (fn phi => fn jset => fn Jz => fn Jz' => - SOME (Proof_Context.cterm_of lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y', + SOME (Thm.cterm_of lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y', HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz)))))) phis jsets Jzs Jzs'; in @@ -2024,7 +2024,7 @@ val goals = @{map 3} mk_goal fs colss colss'; val ctss = - map (fn phi => map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' phi, nat]) goals; + map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) goals; val thms = @{map 4} (fn goal => fn cts => fn rec_0s => fn rec_Sucs => @@ -2047,7 +2047,7 @@ val goals = map (mk_goal Jbds) colss; - val ctss = map (fn phi => map (SOME o Proof_Context.cterm_of lthy) [Term.absfree nat' phi, nat]) + val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) (map (mk_goal (replicate n sbd)) colss); val thms = @@ -2064,7 +2064,7 @@ val map_cong0_thms = let - val cTs = map (SOME o Proof_Context.ctyp_of lthy o + val cTs = map (SOME o Thm.ctyp_of lthy o Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params; fun mk_prem z set f g y y' = @@ -2086,7 +2086,7 @@ HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy) |> Term.absfree y'_copy |> Term.absfree y' - |> Proof_Context.cterm_of lthy; + |> Thm.cterm_of lthy; val cphis = @{map 9} mk_cphi Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy; @@ -2189,9 +2189,9 @@ activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds; fun mk_cts zs z's phis = @{map 3} (fn z => fn z' => fn phi => - SOME (Proof_Context.cterm_of lthy (fold_rev (Term.absfree o Term.dest_Free) [z', z] phi))) + SOME (Thm.cterm_of lthy (fold_rev (Term.absfree o Term.dest_Free) [z', z] phi))) zs z's phis @ - map (SOME o Proof_Context.cterm_of lthy) (splice z's zs); + map (SOME o Thm.cterm_of lthy) (splice z's zs); val cts1 = mk_cts Jzs Jzs_copy coind1_phis; val cts2 = mk_cts Jz's Jz's_copy coind2_phis; @@ -2228,9 +2228,9 @@ Jphis abs fstABs sndABs; val ctss = map2 (fn ab' => fn phis => map2 (fn z' => fn phi => - SOME (Proof_Context.cterm_of lthy (Term.absfree ab' (Term.absfree z' phi)))) + SOME (Thm.cterm_of lthy (Term.absfree ab' (Term.absfree z' phi)))) zip_zs' phis @ - map (SOME o Proof_Context.cterm_of lthy) zip_zs) + map (SOME o Thm.cterm_of lthy) zip_zs) abs' helper_ind_phiss; fun mk_helper_ind_concl ab' z ind_phi set = mk_Ball (set $ z) (Term.absfree ab' ind_phi); diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Fri Mar 06 15:58:56 2015 +0100 @@ -43,7 +43,7 @@ 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))); + (Thm.cterm_of ctxt (Var f), Thm.cterm_of ctxt (mk_proj T num_frees (find s))); val cfps = map mk_cfp fs; in Drule.cterm_instantiate cfps thm @@ -155,7 +155,7 @@ let val s = Name.uu; val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0)); - val split' = Drule.instantiate' [] [SOME (Proof_Context.cterm_of ctxt eq)] split; + val split' = Drule.instantiate' [] [SOME (Thm.cterm_of ctxt eq)] split; in Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split' end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_lfp.ML --- a/src/HOL/Tools/BNF/bnf_lfp.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp.ML Fri Mar 06 15:58:56 2015 +0100 @@ -607,8 +607,8 @@ fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd; val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss); - val card_cT = Proof_Context.ctyp_of lthy suc_bdT; - val card_ct = Proof_Context.cterm_of lthy (Term.absfree idx' card_conjunction); + val card_cT = Thm.ctyp_of lthy suc_bdT; + val card_ct = Thm.cterm_of lthy (Term.absfree idx' card_conjunction); val card_of = Goal.prove_sorry lthy [] [] @@ -622,8 +622,8 @@ val least_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs); - val least_cT = Proof_Context.ctyp_of lthy suc_bdT; - val least_ct = Proof_Context.cterm_of lthy (Term.absfree idx' least_conjunction); + val least_cT = Thm.ctyp_of lthy suc_bdT; + val least_ct = Thm.cterm_of lthy (Term.absfree idx' least_conjunction); val least = (Goal.prove_sorry lthy [] [] @@ -779,7 +779,7 @@ val car_inits = map (mk_min_alg str_inits) ks; - val alg_init_thm = cterm_instantiate_pos (map (SOME o Proof_Context.cterm_of lthy) str_inits) alg_min_alg_thm; + val alg_init_thm = cterm_instantiate_pos (map (SOME o Thm.cterm_of lthy) str_inits) alg_min_alg_thm; val alg_select_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_Ball II @@ -812,7 +812,7 @@ fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_fun_eq init_fs init_fs_copy init_xs)); - val cts = map (Proof_Context.cterm_of lthy) ss; + val cts = map (Thm.cterm_of lthy) ss; val unique_mor = Goal.prove_sorry lthy [] [] (Logic.list_implies (prems @ mor_prems, unique)) (K (mk_init_unique_mor_tac cts m alg_def alg_init_thm least_min_alg_thms @@ -946,7 +946,7 @@ |> Thm.close_derivation; fun mk_ct initFT str abs = Term.absdummy initFT (abs $ (str $ Bound 0)) - val cts = @{map 3} (Proof_Context.cterm_of lthy ooo mk_ct) init_FTs str_inits Abs_Ts; + val cts = @{map 3} (Thm.cterm_of lthy ooo mk_ct) init_FTs str_inits Abs_Ts; val mor_Abs = Goal.prove_sorry lthy [] [] @@ -1020,8 +1020,8 @@ val mor_fold_thm = let val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks); - val cT = Proof_Context.ctyp_of lthy foldT; - val ct = Proof_Context.cterm_of lthy fold_fun + val cT = Thm.ctyp_of lthy foldT; + val ct = Thm.cterm_of lthy fold_fun in Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks))) @@ -1244,7 +1244,7 @@ rev (Term.add_tfrees goal [])) end; - val cTs = map (SOME o Proof_Context.ctyp_of lthy o TFree) induct_params; + val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct_params; val weak_ctor_induct_thms = let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI); @@ -1276,7 +1276,7 @@ (@{map 3} mk_concl phi2s Izs1 Izs2)); fun mk_t phi (z1, z1') (z2, z2') = Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2)); - val cts = @{map 3} (SOME o Proof_Context.cterm_of lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2'); + val cts = @{map 3} (SOME o Thm.cterm_of lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2'); val goal = Logic.list_implies (prems, concl); in (Goal.prove_sorry lthy [] [] goal @@ -1552,7 +1552,7 @@ val timer = time (timer "set functions for the new datatypes"); - val cxs = map (SOME o Proof_Context.cterm_of lthy) Izs; + val cxs = map (SOME o Thm.cterm_of lthy) Izs; val Isetss_by_range' = map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range; @@ -1561,10 +1561,10 @@ fun mk_set_map0 f map z set set' = HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z)); - fun mk_cphi f map z set set' = Proof_Context.cterm_of lthy + fun mk_cphi f map z set set' = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_map0 f map z set set')); - val csetss = map (map (Proof_Context.cterm_of lthy)) Isetss_by_range'; + val csetss = map (map (Thm.cterm_of lthy)) Isetss_by_range'; val cphiss = @{map 3} (fn f => fn sets => fn sets' => (@{map 4} (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range'; @@ -1594,7 +1594,7 @@ let fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd; - fun mk_cphi z set = Proof_Context.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_bd z sbd0 set)); + fun mk_cphi z set = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_bd z sbd0 set)); val cphiss = map (map2 mk_cphi Izs) Isetss_by_range; @@ -1630,7 +1630,7 @@ HOLogic.mk_eq (fmap $ z, gmap $ z)); fun mk_cphi sets z fmap gmap = - Proof_Context.cterm_of lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap)); + Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap)); val cphis = @{map 4} mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps; @@ -1691,9 +1691,9 @@ Irelpsi12 $ Iz1 $ Iz2); val goals = @{map 5} mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2; - val cTs = map (SOME o Proof_Context.ctyp_of lthy o TFree) induct2_params; - val cxs = map (SOME o Proof_Context.cterm_of lthy) (splice Izs1 Izs2); - fun mk_cphi z1 z2 goal = SOME (Proof_Context.cterm_of lthy (Term.absfree z1 (Term.absfree z2 goal))); + val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct2_params; + val cxs = map (SOME o Thm.cterm_of lthy) (splice Izs1 Izs2); + fun mk_cphi z1 z2 goal = SOME (Thm.cterm_of lthy (Term.absfree z1 (Term.absfree z2 goal))); val cphis = @{map 3} mk_cphi Izs1' Izs2' goals; val induct = Drule.instantiate' cTs (cphis @ cxs) ctor_induct2_thm; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML Fri Mar 06 15:58:56 2015 +0100 @@ -273,7 +273,7 @@ HOLogic.mk_Trueprop (BNF_LFP_Util.mk_not_eq (list_comb (size, xs)) HOLogic.zero); val thm = Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => - mk_size_neq ctxt (map (Proof_Context.cterm_of lthy2) xs) + mk_size_neq ctxt (map (Thm.cterm_of lthy2) xs) (#exhaust (#ctr_sugar (#fp_ctr_sugar fp_sugar))) size_thms) |> single |> Proof_Context.export names_lthy2 lthy2 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML Fri Mar 06 15:58:56 2015 +0100 @@ -679,10 +679,10 @@ let val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; - val rho_As = map (apply2 (Proof_Context.ctyp_of lthy)) (map Logic.varifyT_global As ~~ As); + val rho_As = map (apply2 (Thm.ctyp_of lthy)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = - Drule.instantiate' [] [SOME (Proof_Context.cterm_of lthy t)] + Drule.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; @@ -997,7 +997,7 @@ mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; in Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals) - (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Proof_Context.cterm_of ctxt u) + (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Conjunction.elim_balanced (length goals) |> Proof_Context.export names_lthy lthy @@ -1020,7 +1020,7 @@ val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => - mk_case_distrib_tac ctxt (Proof_Context.cterm_of ctxt u) exhaust_thm case_thms) + mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Fri Mar 06 15:58:56 2015 +0100 @@ -36,7 +36,7 @@ |> Thm.prop_of |> Logic.dest_equals |> fst |> Term.strip_comb ||> fst o split_last |> list_comb; val lhs = Free (singleton (Name.variant_list params) "case", Term.fastype_of rhs); - val assum = Thm.cterm_of thy (Logic.mk_equals (lhs, rhs)); + val assum = Thm.global_cterm_of thy (Logic.mk_equals (lhs, rhs)); in thms |> Conjunction.intr_balanced diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML Fri Mar 06 15:58:56 2015 +0100 @@ -219,7 +219,7 @@ val ps = map_filter (fn (_, NONE) => NONE - | (var, SOME ct) => SOME (Thm.cterm_of thy (Var var), ct)) (vars' ~~ cts); + | (var, SOME ct) => SOME (Thm.global_cterm_of thy (Var var), ct)) (vars' ~~ cts); in Drule.cterm_instantiate ps thm end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/fun_cases.ML --- a/src/HOL/Tools/Function/fun_cases.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/fun_cases.ML Fri Mar 06 15:58:56 2015 +0100 @@ -29,7 +29,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 = Proof_Context.cterm_of ctxt prop; + val cprop = Thm.cterm_of ctxt prop; fun mk_elim rl = Thm.implies_intr cprop (Tactic.rule_by_tactic ctxt (Inductive.mk_cases_tac ctxt) (Thm.assume cprop RS rl)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/function_common.ML --- a/src/HOL/Tools/Function/function_common.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/function_common.ML Fri Mar 06 15:58:56 2015 +0100 @@ -327,7 +327,7 @@ fun lift_morphism ctxt f = let - fun term t = Thm.term_of (Drule.cterm_rule f (Proof_Context.cterm_of ctxt t)) + fun term t = Thm.term_of (Drule.cterm_rule f (Thm.cterm_of ctxt t)) in Morphism.morphism "lift_morphism" {binding = [], @@ -338,11 +338,11 @@ fun import_function_data t ctxt = let - val ct = Proof_Context.cterm_of ctxt t + val ct = Thm.cterm_of ctxt t val inst_morph = lift_morphism ctxt o Thm.instantiate fun match (trm, data) = - SOME (transform_function_data data (inst_morph (Thm.match (Proof_Context.cterm_of ctxt trm, ct)))) + SOME (transform_function_data data (inst_morph (Thm.match (Thm.cterm_of ctxt trm, ct)))) handle Pattern.MATCH => NONE in get_first match (Item_Net.retrieve (get_functions ctxt) t) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/function_context_tree.ML --- a/src/HOL/Tools/Function/function_context_tree.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/function_context_tree.ML Fri Mar 06 15:58:56 2015 +0100 @@ -119,7 +119,7 @@ val subst = Pattern.match thy (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 => apply2 (Proof_Context.cterm_of ctxt) (Var v, Envir.subst_term subst (Var v))) + map (fn v => apply2 (Thm.cterm_of ctxt) (Var v, Envir.subst_term subst (Var v))) (Term.add_vars c []) in (cterm_instantiate inst r, dep, branches) @@ -147,7 +147,7 @@ val (r, dep, branches) = find_cong_rule ctxt fvar h congs_deps t fun subtree (ctxt', fixes, assumes, st) = ((fixes, - map (Thm.assume o Proof_Context.cterm_of ctxt) assumes), + map (Thm.assume o Thm.cterm_of ctxt) assumes), mk_tree' ctxt' st) in Cong (r, dep, map subtree branches) @@ -158,8 +158,8 @@ fun inst_tree ctxt fvar f tr = let - val cfvar = Proof_Context.cterm_of ctxt fvar - val cf = Proof_Context.cterm_of ctxt f + val cfvar = Thm.cterm_of ctxt fvar + val cf = Thm.cterm_of ctxt f fun inst_term t = subst_bound(f, abstract_over (fvar, t)) @@ -172,7 +172,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 Proof_Context.cterm_of ctxt o inst_term o Thm.prop_of) assms), + ((fxs, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) assms), inst_tree_aux str) in inst_tree_aux tr @@ -188,10 +188,10 @@ fun export_thm ctxt (fixes, assumes) = fold_rev (Thm.implies_intr o Thm.cprop_of) assumes - #> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt o Free) fixes + #> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) fixes fun import_thm ctxt (fixes, athms) = - fold (Thm.forall_elim o Proof_Context.cterm_of ctxt o Free) fixes + fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) fixes #> fold Thm.elim_implies athms @@ -240,7 +240,7 @@ fun rewrite_by_tree ctxt h ih x tr = let - fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (Proof_Context.cterm_of ctxt t), x) + fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (Thm.cterm_of ctxt 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" *) @@ -248,10 +248,10 @@ val iha = import_thm ctxt (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 (Proof_Context.cterm_of ctxt arg)] ih + val inst_ih = instantiate' [] [SOME (Thm.cterm_of ctxt 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 (Proof_Context.cterm_of ctxt h)) inner + val h_a'_eq_h_a = Thm.combination (Thm.reflexive (Thm.cterm_of ctxt 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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/function_core.ML Fri Mar 06 15:58:56 2015 +0100 @@ -150,11 +150,11 @@ val lhs = inst pre_lhs val rhs = inst pre_rhs - val cqs = map (Proof_Context.cterm_of ctxt') qs - val ags = map (Thm.assume o Proof_Context.cterm_of ctxt') gs + val cqs = map (Thm.cterm_of ctxt') qs + val ags = map (Thm.assume o Thm.cterm_of ctxt') gs val case_hyp = - Thm.assume (Proof_Context.cterm_of ctxt' (HOLogic.mk_Trueprop (mk_eq (x, lhs)))) + Thm.assume (Thm.cterm_of ctxt' (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 } @@ -185,7 +185,7 @@ (* Instantiate the GIntro thm with "f" and import into the clause context. *) val lGI = GIntro_thm - |> Thm.forall_elim (Proof_Context.cterm_of ctxt f) + |> Thm.forall_elim (Thm.cterm_of ctxt f) |> fold Thm.forall_elim cqs |> fold Thm.elim_implies ags @@ -193,7 +193,7 @@ let val llRI = RI |> fold Thm.forall_elim cqs - |> fold (Thm.forall_elim o Proof_Context.cterm_of ctxt o Free) rcfix + |> fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) rcfix |> fold Thm.elim_implies ags |> fold Thm.elim_implies rcassm @@ -233,7 +233,7 @@ val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj - val lhsi_eq_lhsj = Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj))) + val lhsi_eq_lhsj = Thm.global_cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj))) in if j < i then let val compat = lookup_compat_thm j i cts @@ -271,7 +271,7 @@ val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs val h_assums = map (fn RCInfo {h_assum, ...} => - Thm.assume (Proof_Context.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs + Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs val (eql, _) = Function_Context_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree @@ -301,15 +301,15 @@ val compat = get_compat_thm thy compat_store i j cctxi cctxj val Ghsj' = map (fn RCInfo {h_assum, ...} => - Thm.assume (Thm.cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj + Thm.assume (Thm.global_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 (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')))) + val y_eq_rhsj'h = Thm.assume (Thm.global_cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h)))) + val lhsi_eq_lhsj' = Thm.assume (Thm.global_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' *) @@ -324,7 +324,7 @@ (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *) |> 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') + |> fold_rev Thm.forall_intr (Thm.global_cterm_of thy h :: cqsj') end @@ -340,12 +340,12 @@ fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case) |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt o Free) RIvs + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs val existence = fold (curry op COMP o prep_RC) RCs lGI - val P = Proof_Context.cterm_of ctxt (mk_eq (y, rhsC)) - val G_lhs_y = Thm.assume (Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y))) + val P = Thm.cterm_of ctxt (mk_eq (y, rhsC)) + val G_lhs_y = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y))) val unique_clauses = map2 (mk_uniqueness_clause thy globals compat_store clausei) clauses rep_lemmas @@ -356,21 +356,21 @@ |> Seq.list_of |> the_single val uniqueness = G_cases - |> Thm.forall_elim (Proof_Context.cterm_of ctxt lhs) - |> Thm.forall_elim (Proof_Context.cterm_of ctxt y) + |> Thm.forall_elim (Thm.cterm_of ctxt lhs) + |> Thm.forall_elim (Thm.cterm_of ctxt y) |> Thm.forall_elim P |> Thm.elim_implies G_lhs_y |> fold elim_implies_eta unique_clauses |> Thm.implies_intr (Thm.cprop_of G_lhs_y) - |> Thm.forall_intr (Proof_Context.cterm_of ctxt y) + |> Thm.forall_intr (Thm.cterm_of ctxt y) - val P2 = Proof_Context.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *) + val P2 = Thm.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *) val exactly_one = @{thm ex1I} |> instantiate' - [SOME (Proof_Context.ctyp_of ctxt ranT)] - [SOME P2, SOME (Proof_Context.cterm_of ctxt rhsC)] + [SOME (Thm.ctyp_of ctxt ranT)] + [SOME P2, SOME (Thm.cterm_of ctxt rhsC)] |> curry (op COMP) existence |> curry (op COMP) uniqueness |> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym]) @@ -382,8 +382,8 @@ existence |> Thm.implies_intr ihyp |> Thm.implies_intr (Thm.cprop_of case_hyp) - |> Thm.forall_intr (Proof_Context.cterm_of ctxt x) - |> Thm.forall_elim (Proof_Context.cterm_of ctxt lhs) + |> Thm.forall_intr (Thm.cterm_of ctxt x) + |> Thm.forall_elim (Thm.cterm_of ctxt lhs) |> curry (op RS) refl in (exactly_one, function_value) @@ -399,12 +399,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)))) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt 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 (Proof_Context.cterm_of ctxt h)] + |> instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)] val _ = trace_msg (K "Proving Replacement lemmas...") val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses @@ -421,11 +421,11 @@ |> Thm.forall_elim_vars 0 |> fold (curry op COMP) ex1s |> Thm.implies_intr (ihyp) - |> Thm.implies_intr (Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x))) - |> Thm.forall_intr (Proof_Context.cterm_of ctxt x) + |> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x))) + |> Thm.forall_intr (Thm.cterm_of ctxt x) |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *) |> (fn it => - fold (Thm.forall_intr o Proof_Context.cterm_of ctxt o Var) + fold (Thm.forall_intr o Thm.cterm_of ctxt o Var) (Term.add_vars (Thm.prop_of it) []) it) val goalstate = Conjunction.intr graph_is_function complete @@ -466,7 +466,7 @@ #> the_default ("", 0) in fold_rev (fn Free (n, T) => - forall_intr_rename (n, Proof_Context.cterm_of lthy (Var (varmap (n, T), T)))) qs thm + forall_intr_rename (n, Thm.cterm_of lthy (Var (varmap (n, T), T)))) qs thm end in ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy) @@ -552,7 +552,7 @@ let fun inst_term t = subst_bound(f, abstract_over (fvar, t)) in - (rcfix, map (Thm.assume o Thm.cterm_of thy o inst_term o Thm.prop_of) rcassm, inst_term rcarg) + (rcfix, map (Thm.assume o Thm.global_cterm_of thy o inst_term o Thm.prop_of) rcassm, inst_term rcarg) end @@ -569,14 +569,14 @@ (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm = let val lhs_acc = - Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *) + Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *) val z_smaller = - Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *) + Thm.cterm_of ctxt (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 (Proof_Context.cterm_of ctxt z) + |> Thm.forall_intr (Thm.cterm_of ctxt z) |> (fn it => it COMP valthm) |> Thm.implies_intr lhs_acc |> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff]) @@ -599,23 +599,23 @@ val Globals {domT, x, z, a, P, D, ...} = globals val acc_R = mk_acc domT R - 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 x_D = Thm.assume (Thm.global_cterm_of thy (HOLogic.mk_Trueprop (D $ x))) + val a_D = Thm.global_cterm_of thy (HOLogic.mk_Trueprop (D $ a)) - val D_subset = Thm.cterm_of thy (Logic.all x + val D_subset = Thm.global_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))))) - |> Thm.cterm_of thy + |> Thm.global_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))) - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy val aihyp = Thm.assume ihyp @@ -633,10 +633,10 @@ end fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih - |> Thm.forall_elim (Thm.cterm_of thy rcarg) + |> Thm.forall_elim (Thm.global_cterm_of thy rcarg) |> Thm.elim_implies llRI |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas - |> fold_rev (Thm.forall_intr o Thm.cterm_of thy o Free) RIvs + |> fold_rev (Thm.forall_intr o Thm.global_cterm_of thy o Free) RIvs val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *) @@ -645,7 +645,7 @@ |> 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) - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy val P_lhs = Thm.assume step |> fold Thm.forall_elim cqs @@ -653,7 +653,7 @@ |> fold Thm.elim_implies ags |> fold Thm.elim_implies P_recs - val res = Thm.cterm_of thy (HOLogic.mk_Trueprop (P $ x)) + val res = Thm.global_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" *) @@ -671,7 +671,7 @@ |> fold (curry op COMP) cases (* P x *) |> Thm.implies_intr ihyp |> Thm.implies_intr (Thm.cprop_of x_D) - |> Thm.forall_intr (Thm.cterm_of thy x) + |> Thm.forall_intr (Thm.global_cterm_of thy x) val subset_induct_rule = acc_subset_induct @@ -686,16 +686,16 @@ val simple_induct_rule = subset_induct_rule - |> Thm.forall_intr (Thm.cterm_of thy D) - |> Thm.forall_elim (Thm.cterm_of thy acc_R) + |> Thm.forall_intr (Thm.global_cterm_of thy D) + |> Thm.forall_elim (Thm.global_cterm_of thy acc_R) |> atac 1 |> Seq.hd |> (curry op COMP) (acc_downward - |> (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) + |> (instantiate' [SOME (Thm.global_ctyp_of thy domT)] + (map (SOME o Thm.global_cterm_of thy) [R, x, z])) + |> Thm.forall_intr (Thm.global_cterm_of thy z) + |> Thm.forall_intr (Thm.global_cterm_of thy x)) + |> Thm.forall_intr (Thm.global_cterm_of thy a) + |> Thm.forall_intr (Thm.global_cterm_of thy P) in simple_induct_rule end @@ -708,7 +708,7 @@ qglr = (oqs, _, _, _), ...} = clause val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs) |> fold_rev (curry Logic.mk_implies) gs - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt in Goal.init goal |> (SINGLE (resolve_tac ctxt [accI] 1)) |> the @@ -744,7 +744,7 @@ |> Function_Context_Tree.export_term (fixes, assumes) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags |> fold_rev mk_forall_rename (map fst oqs ~~ qs) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val thm = Thm.assume hyp |> fold Thm.forall_elim cqs @@ -753,7 +753,7 @@ |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *) val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg)) - |> Proof_Context.cterm_of ctxt |> Thm.assume + |> Thm.cterm_of ctxt |> Thm.assume val acc = thm COMP ih_case val z_acc_local = acc @@ -788,42 +788,42 @@ val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R') - |> Proof_Context.cterm_of ctxt (* "wf R'" *) + |> Thm.cterm_of ctxt (* "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))) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0 - val R_z_x = Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ x)) + val R_z_x = Thm.cterm_of ctxt (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 (Proof_Context.cterm_of ctxt z) - |> Thm.forall_elim (Proof_Context.cterm_of ctxt x) - |> Thm.forall_elim (Proof_Context.cterm_of ctxt (acc_R $ z)) + |> Thm.forall_elim (Thm.cterm_of ctxt z) + |> Thm.forall_elim (Thm.cterm_of ctxt x) + |> Thm.forall_elim (Thm.cterm_of ctxt (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 (Proof_Context.cterm_of ctxt z) + |> Thm.forall_intr (Thm.cterm_of ctxt z) |> (fn it => it COMP accI) |> Thm.implies_intr ihyp - |> Thm.forall_intr (Proof_Context.cterm_of ctxt x) + |> Thm.forall_intr (Thm.cterm_of ctxt 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 (Proof_Context.cterm_of ctxt R') - |> Thm.forall_elim (Proof_Context.cterm_of ctxt (inrel_R)) + |> Thm.forall_intr (Thm.cterm_of ctxt R') + |> Thm.forall_elim (Thm.cterm_of ctxt (inrel_R)) |> curry op RS wf_in_rel |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps [in_rel_def]) - |> Thm.forall_intr (Proof_Context.cterm_of ctxt Rrel) + |> Thm.forall_intr (Thm.cterm_of ctxt Rrel) end @@ -878,11 +878,11 @@ RCss GIntro_thms) RIntro_thmss val complete = - mk_completeness globals clauses abstract_qglrs |> Proof_Context.cterm_of lthy |> Thm.assume + mk_completeness globals clauses abstract_qglrs |> Thm.cterm_of lthy |> Thm.assume val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs - |> map (Proof_Context.cterm_of lthy #> Thm.assume) + |> map (Thm.cterm_of lthy #> Thm.assume) val compat_store = store_compat_thms n compat diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/function_elims.ML --- a/src/HOL/Tools/Function/function_elims.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/function_elims.ML Fri Mar 06 15:58:56 2015 +0100 @@ -115,12 +115,12 @@ val args = HOLogic.mk_tuple arg_vars; val domT = R |> dest_Free |> snd |> hd o snd o dest_Type; - val P = mk_var "P" @{typ "bool"} name_ctxt |> fst |> Proof_Context.cterm_of ctxt + val P = mk_var "P" @{typ "bool"} name_ctxt |> fst |> Thm.cterm_of ctxt val sumtree_inj = Sum_Tree.mk_inj domT n_fs (idx+1) args; - val cprop = Proof_Context.cterm_of ctxt prop; + val cprop = Thm.cterm_of ctxt prop; - val asms = [cprop, Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (dom $ sumtree_inj))]; + val asms = [cprop, Thm.cterm_of ctxt (HOLogic.mk_Trueprop (dom $ sumtree_inj))]; val asms_thms = map Thm.assume asms; fun prep_subgoal_tac i = @@ -133,10 +133,10 @@ val elim_stripped = nth cases idx |> Thm.forall_elim P - |> Thm.forall_elim (Proof_Context.cterm_of ctxt args) + |> Thm.forall_elim (Thm.cterm_of ctxt args) |> Tactic.rule_by_tactic ctxt (ALLGOALS prep_subgoal_tac) |> fold_rev Thm.implies_intr asms - |> Thm.forall_intr (Proof_Context.cterm_of ctxt rhs_var); + |> Thm.forall_intr (Thm.cterm_of ctxt rhs_var); val bool_elims = (case ranT of @@ -145,7 +145,7 @@ fun unstrip rl = rl - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt) arg_vars + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) arg_vars |> Thm.forall_intr P in map unstrip (elim_stripped :: bool_elims) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/function_lib.ML --- a/src/HOL/Tools/Function/function_lib.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/function_lib.ML Fri Mar 06 15:58:56 2015 +0100 @@ -105,7 +105,7 @@ val ty = fastype_of t in Goal.prove_internal ctxt [] - (Proof_Context.cterm_of ctxt + (Thm.cterm_of ctxt (Logic.mk_equals (t, if null is then mk (Const (neu, ty), foldr1 mk (map (nth xs) js)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/induction_schema.ML --- a/src/HOL/Tools/Function/induction_schema.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/induction_schema.ML Fri Mar 06 15:58:56 2015 +0100 @@ -46,7 +46,7 @@ (map meta (@{thm split_conv} :: @{thms sum.case})) fun term_conv thy cv t = - cv (Thm.cterm_of thy t) + cv (Thm.global_cterm_of thy t) |> Thm.prop_of |> Logic.dest_equals |> snd fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)) @@ -221,7 +221,7 @@ $ (HOLogic.pair_const T T $ Bound 0 $ x) $ R), HOLogic.mk_Trueprop (P_comp $ Bound 0))) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val aihyp = Thm.assume ihyp @@ -235,9 +235,9 @@ let val fxs = map Free xs val branch_hyp = - Thm.assume (Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat)))) + Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat)))) - val C_hyps = map (Proof_Context.cterm_of ctxt #> Thm.assume) Cs + val C_hyps = map (Thm.cterm_of ctxt #> Thm.assume) Cs val (relevant_cases, ineqss') = (scases_idx ~~ ineqss) @@ -247,11 +247,11 @@ fun prove_case (cidx, SchemeCase {qs, gs, lhs, rs, ...}) ineq_press = let val case_hyps = - map (Thm.assume o Proof_Context.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) + map (Thm.assume o Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs) - val cqs = map (Proof_Context.cterm_of ctxt o Free) qs - val ags = map (Thm.assume o Proof_Context.cterm_of ctxt) gs + val cqs = map (Thm.cterm_of ctxt o Free) qs + val ags = map (Thm.assume o Thm.cterm_of ctxt) gs val replace_x_simpset = put_simpset HOL_basic_ss ctxt addsimps (branch_hyp :: case_hyps) @@ -259,16 +259,16 @@ fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) = let - val cGas = map (Thm.assume o Proof_Context.cterm_of ctxt) Gas - val cGvs = map (Proof_Context.cterm_of ctxt o Free) Gvs + val cGas = map (Thm.assume o Thm.cterm_of ctxt) Gas + val cGvs = map (Thm.cterm_of ctxt o Free) Gvs val import = fold Thm.forall_elim (cqs @ cGvs) #> fold Thm.elim_implies (ags @ cGas) val ipres = pres - |> Thm.forall_elim (Proof_Context.cterm_of ctxt (list_comb (P_of idx, rcargs))) + |> Thm.forall_elim (Thm.cterm_of ctxt (list_comb (P_of idx, rcargs))) |> import in sih - |> Thm.forall_elim (Proof_Context.cterm_of ctxt (inject idx rcargs)) + |> Thm.forall_elim (Thm.cterm_of ctxt (inject idx rcargs)) |> Thm.elim_implies (import ineq) (* Psum rcargs *) |> Conv.fconv_rule (sum_prod_conv ctxt) |> Conv.fconv_rule (ind_rulify ctxt) @@ -283,7 +283,7 @@ |> 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 - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val Plhs_to_Pxs_conv = foldl1 (uncurry Conv.combination_conv) @@ -303,15 +303,15 @@ val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss') val bstep = complete_thm - |> Thm.forall_elim (Proof_Context.cterm_of ctxt (list_comb (P, fxs))) - |> fold (Thm.forall_elim o Proof_Context.cterm_of ctxt) (fxs @ map Free ws) + |> Thm.forall_elim (Thm.cterm_of ctxt (list_comb (P, fxs))) + |> fold (Thm.forall_elim o Thm.cterm_of ctxt) (fxs @ map Free ws) |> fold Thm.elim_implies C_hyps |> fold Thm.elim_implies cases (* P xs *) |> fold_rev (Thm.implies_intr o Thm.cprop_of) C_hyps - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt o Free) ws + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) ws val Pxs = - Proof_Context.cterm_of ctxt (HOLogic.mk_Trueprop (P_comp $ x)) + Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P_comp $ x)) |> Goal.init |> (Simplifier.rewrite_goals_tac ctxt (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.case})) @@ -320,7 +320,7 @@ |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep) |> Goal.finish ctxt |> Thm.implies_intr (Thm.cprop_of branch_hyp) - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt) fxs + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) fxs in (Pxs, steps) end @@ -332,7 +332,7 @@ val istep = sum_split_rule |> fold (fn b => fn th => Drule.compose (b, 1, th)) branches |> Thm.implies_intr ihyp - |> Thm.forall_intr (Proof_Context.cterm_of ctxt x) (* "!!x. (!!y P x" *) + |> Thm.forall_intr (Thm.cterm_of ctxt x) (* "!!x. (!!y P x" *) val induct_rule = @{thm "wf_induct_rule"} @@ -357,11 +357,11 @@ val ineqss = mk_ineqs R scheme - |> map (map (apply2 (Thm.assume o Proof_Context.cterm_of ctxt''))) + |> map (map (apply2 (Thm.assume o Thm.cterm_of ctxt''))) val complete = - map_range (mk_completeness ctxt'' scheme #> Proof_Context.cterm_of ctxt'' #> Thm.assume) + map_range (mk_completeness ctxt'' scheme #> Thm.cterm_of ctxt'' #> Thm.assume) (length branches) - val wf_thm = mk_wf R scheme |> Proof_Context.cterm_of ctxt'' |> Thm.assume + val wf_thm = mk_wf R scheme |> Thm.cterm_of ctxt'' |> Thm.assume val (descent, pres) = split_list (flat ineqss) val newgoals = complete @ pres @ wf_thm :: descent @@ -373,7 +373,7 @@ let val inst = (foldr1 HOLogic.mk_prod (map Free xs)) |> Sum_Tree.mk_inj ST (length branches) (i + 1) - |> Proof_Context.cterm_of ctxt'' + |> Thm.cterm_of ctxt'' in indthm |> Drule.instantiate' [] [SOME inst] diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Fri Mar 06 15:58:56 2015 +0100 @@ -68,7 +68,7 @@ fun mk_cell ctxt solve_tac (vars, prems, lhs, rhs) mfun = Lazy.lazy (fn _ => let - val goals = Proof_Context.cterm_of ctxt o mk_goal (vars, prems, mfun $ lhs, mfun $ rhs) + val goals = Thm.cterm_of ctxt 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 @@ -201,7 +201,7 @@ in (* 4: proof reconstruction *) st |> - (PRIMITIVE (cterm_instantiate [apply2 (Proof_Context.cterm_of ctxt) (rel, relation)]) + (PRIMITIVE (cterm_instantiate [apply2 (Thm.cterm_of ctxt) (rel, relation)]) THEN (REPEAT (rtac @{thm "wf_mlex"} 1)) THEN (rtac @{thm "wf_empty"} 1) THEN EVERY (map prove_row clean_table)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/measure_functions.ML --- a/src/HOL/Tools/Function/measure_functions.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/measure_functions.ML Fri Mar 06 15:58:56 2015 +0100 @@ -20,7 +20,7 @@ fun find_measures ctxt T = 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) + |> Thm.cterm_of ctxt |> Goal.init) |> Seq.map (Thm.prop_of #> (fn _ $ (_ $ (_ $ f)) => f)) |> Seq.list_of diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/mutual.ML Fri Mar 06 15:58:56 2015 +0100 @@ -158,8 +158,8 @@ val args = map inst pre_args val rhs = inst pre_rhs - val cqs = map (Proof_Context.cterm_of ctxt) qs - val ags = map (Thm.assume o Proof_Context.cterm_of ctxt) gs + val cqs = map (Thm.cterm_of ctxt) qs + val ags = map (Thm.assume o Thm.cterm_of ctxt) gs val import = fold Thm.forall_elim cqs #> fold Thm.elim_implies ags @@ -198,7 +198,7 @@ let val xs = map_index (fn (i, T) => - Proof_Context.cterm_of ctxt + Thm.cterm_of ctxt (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 @@ -226,7 +226,7 @@ val case_exp = Sum_Tree.mk_sumcases HOLogic.boolT Ps val induct_inst = - Thm.forall_elim (Proof_Context.cterm_of ctxt case_exp) induct + Thm.forall_elim (Thm.cterm_of ctxt case_exp) induct |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps all_f_defs) @@ -236,9 +236,9 @@ val inj = Sum_Tree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs) in (rule - |> Thm.forall_elim (Proof_Context.cterm_of ctxt inj) + |> Thm.forall_elim (Thm.cterm_of ctxt inj) |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt) - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt) (afs @ newPs), + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) (afs @ newPs), k + length cargTs) end in @@ -258,7 +258,7 @@ val argsT = fastype_of (HOLogic.mk_tuple arg_vars) val (args, name_ctxt) = mk_var "x" argsT name_ctxt - val P = mk_var "P" @{typ "bool"} name_ctxt |> fst |> Proof_Context.cterm_of ctxt + val P = mk_var "P" @{typ "bool"} name_ctxt |> fst |> Thm.cterm_of ctxt val sumtree_inj = Sum_Tree.mk_inj ST n i args val sum_elims = @@ -270,9 +270,9 @@ in cases_rule |> Thm.forall_elim P - |> Thm.forall_elim (Proof_Context.cterm_of ctxt sumtree_inj) + |> Thm.forall_elim (Thm.cterm_of ctxt sumtree_inj) |> Tactic.rule_by_tactic ctxt (ALLGOALS prep_subgoal) - |> Thm.forall_intr (Proof_Context.cterm_of ctxt args) + |> Thm.forall_intr (Thm.cterm_of ctxt args) |> Thm.forall_intr P end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/partial_function.ML --- a/src/HOL/Tools/Function/partial_function.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/partial_function.ML Fri Mar 06 15:58:56 2015 +0100 @@ -113,7 +113,7 @@ let val thy = Thm.theory_of_thm thm; val vs = rev (Term.add_vars (Thm.prop_of thm) []) - |> map (Thm.cterm_of thy o Var); + |> map (Thm.global_cterm_of thy o Var); in cterm_instantiate (zip_options vs cts) thm end; @@ -168,7 +168,7 @@ curry induction predicate *) fun specialize_fixp_induct ctxt args fT fT_uc F curry uncurry mono_thm f_def rule = let - val cert = Proof_Context.cterm_of ctxt + val cert = Thm.cterm_of ctxt val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt val P_inst = Abs ("f", fT_uc, Free (P, fT --> HOLogic.boolT) $ (curry $ Bound 0)) in @@ -187,7 +187,7 @@ fun mk_curried_induct args ctxt inst_rule = let - val cert = Proof_Context.cterm_of ctxt + val cert = Thm.cterm_of ctxt val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt val split_paired_all_conv = @@ -234,7 +234,7 @@ val ((f_binding, fT), mixfix) = the_single fixes; val fname = Binding.name_of f_binding; - val cert = Proof_Context.cterm_of lthy; + val cert = Thm.cterm_of lthy; val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn); val (head, args) = strip_comb lhs; val argnames = map (fst o dest_Free) args; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/pat_completeness.ML --- a/src/HOL/Tools/Function/pat_completeness.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/pat_completeness.ML Fri Mar 06 15:58:56 2015 +0100 @@ -26,7 +26,7 @@ fun inst_case_thm thy x P thm = let val [Pv, xv] = Term.add_vars (Thm.prop_of thm) [] in - thm |> cterm_instantiate (map (apply2 (Thm.cterm_of thy)) [(Var xv, x), (Var Pv, P)]) + thm |> cterm_instantiate (map (apply2 (Thm.global_cterm_of thy)) [(Var xv, x), (Var Pv, P)]) end fun invent_vars constr i = @@ -45,7 +45,7 @@ | filter_pats ctxt cons pvars (((pat as Free _) :: pats, thm) :: pts) = let val inst = list_comb (cons, pvars) in (inst :: pats, - inst_free (Proof_Context.cterm_of ctxt pat) (Proof_Context.cterm_of ctxt inst) thm) + inst_free (Thm.cterm_of ctxt pat) (Thm.cterm_of ctxt inst) thm) :: (filter_pats ctxt cons pvars pts) end | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) = @@ -59,7 +59,7 @@ let val (_, subps) = strip_comb pat val eqs = - map (Proof_Context.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps) + map (Thm.cterm_of ctxt 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 @@ -74,14 +74,14 @@ let val (avars, pvars, newidx) = invent_vars cons idx val c_hyp = - Proof_Context.cterm_of ctxt + Thm.cterm_of ctxt (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 ctxt cons pvars pats) in o_alg ctxt P newidx (avars @ vs) newpats |> Thm.implies_intr c_hyp - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt) avars + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) avars end | constr_case _ _ _ _ _ _ = raise Match and o_alg _ P idx [] (([], Pthm) :: _) = Pthm @@ -91,8 +91,8 @@ then o_alg ctxt P idx vs (map (fn (pv :: pats, thm) => (pats, refl RS - (inst_free (Proof_Context.cterm_of ctxt pv) - (Proof_Context.cterm_of ctxt v) thm))) pts) + (inst_free (Thm.cterm_of ctxt pv) + (Thm.cterm_of ctxt v) thm))) pts) else (* Cons case *) let val thy = Proof_Context.theory_of ctxt @@ -112,10 +112,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 - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt val hyps = map2 mk_assum qss patss - fun inst_hyps hyp qs = fold (Thm.forall_elim o Proof_Context.cterm_of ctxt) qs (Thm.assume hyp) + fun inst_hyps hyp qs = fold (Thm.forall_elim o Thm.cterm_of ctxt) qs (Thm.assume hyp) val assums = map2 inst_hyps hyps qss in o_alg ctxt P 2 xs (patss ~~ assums) @@ -142,7 +142,7 @@ val patss = map (map snd) x_pats val complete_thm = prove_completeness ctxt xs thesis qss patss - |> fold_rev (Thm.forall_intr o Proof_Context.cterm_of ctxt) vs + |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) vs in PRIMITIVE (fn st => Drule.compose (complete_thm, i, st)) end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/relation.ML --- a/src/HOL/Tools/Function/relation.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/relation.ML Fri Mar 06 15:58:56 2015 +0100 @@ -19,7 +19,7 @@ (case Term.add_vars (Thm.prop_of st) [] of (* FIXME tactic should not inspect main conclusion *) [v as (_, T)] => let val thy = Thm.theory_of_thm st in - PRIMITIVE (Thm.instantiate ([], [(Thm.cterm_of thy (Var v), Thm.cterm_of thy (inst T))])) st + PRIMITIVE (Thm.instantiate ([], [(Thm.global_cterm_of thy (Var v), Thm.global_cterm_of thy (inst T))])) st end | _ => Seq.empty); @@ -38,9 +38,9 @@ |> map_types Type_Infer.paramify_vars |> Type.constraint T |> Syntax.check_term ctxt - |> Proof_Context.cterm_of ctxt; + |> Thm.cterm_of ctxt; in - PRIMITIVE (Thm.instantiate ([], [(Proof_Context.cterm_of ctxt (Var v), rel')])) st + PRIMITIVE (Thm.instantiate ([], [(Thm.cterm_of ctxt (Var v), rel')])) st end | _ => Seq.empty); diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/scnp_reconstruct.ML --- a/src/HOL/Tools/Function/scnp_reconstruct.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Fri Mar 06 15:58:56 2015 +0100 @@ -272,7 +272,7 @@ val level_mapping = map_index pt_lev lev |> Termination.mk_sumcases D (setT nat_pairT) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt in PROFILE "Proof Reconstruction" (CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv sl))) 1 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Function/termination.ML --- a/src/HOL/Tools/Function/termination.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Function/termination.ML Fri Mar 06 15:58:56 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 (Proof_Context.cterm_of ctxt goal) chain_tac of + case Function_Lib.try_proof (Thm.cterm_of ctxt goal) chain_tac of Function_Lib.Solved thm => SOME thm | _ => NONE end @@ -169,7 +169,7 @@ fun mk_desc ctxt tac vs Gam l r m1 m2 = let fun try rel = - try_proof (Proof_Context.cterm_of ctxt + try_proof (Thm.cterm_of ctxt (Logic.list_all (vs, Logic.mk_implies (HOLogic.mk_Trueprop Gam, HOLogic.mk_Trueprop (Const (rel, @{typ "nat => nat => bool"}) @@ -302,7 +302,7 @@ THEN' (blast_tac ctxt))) (* Solve rest of context... not very elegant *) ) i in - (PRIMITIVE (Drule.cterm_instantiate [apply2 (Proof_Context.cterm_of ctxt) (rel, relation)]) + (PRIMITIVE (Drule.cterm_instantiate [apply2 (Thm.cterm_of ctxt) (rel, relation)]) THEN ALLGOALS (fn i => if i = 1 then all_tac else solve_membership_tac i) THEN rewrite_goal_tac ctxt Un_aci_simps 1) (* eliminate duplicates *) end) 1 st diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Lifting/lifting_bnf.ML --- a/src/HOL/Tools/Lifting/lifting_bnf.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Lifting/lifting_bnf.ML Fri Mar 06 15:58:56 2015 +0100 @@ -23,7 +23,7 @@ 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 + val inst = map2 (curry (apply2 (Thm.cterm_of ctxt))) vars UNIVs val rel_Grp_UNIV_sym = rel_Grp |> Drule.instantiate_normalize ([], inst) |> Local_Defs.unfold ctxt @{thms subset_UNIV[THEN eqTrueI] UNIV_def[symmetric] simp_thms(21)} |> (fn thm => thm RS sym) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Lifting/lifting_def.ML --- a/src/HOL/Tools/Lifting/lifting_def.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Lifting/lifting_def.ML Fri Mar 06 15:58:56 2015 +0100 @@ -57,7 +57,7 @@ fun try_prove_reflexivity ctxt prop = let val thy = Proof_Context.theory_of ctxt - val cprop = Thm.cterm_of thy prop + val cprop = Thm.global_cterm_of thy prop val rule = @{thm ge_eq_refl} val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule) val insts = Thm.first_order_match (concl_pat, cprop) @@ -99,7 +99,7 @@ end; val subst = fold make_subst free_vars []; - val csubst = map (apply2 (Thm.cterm_of thy)) subst; + val csubst = map (apply2 (Thm.global_cterm_of thy)) subst; val inst_thm = Drule.cterm_instantiate csubst thm; in Conv.fconv_rule @@ -112,13 +112,13 @@ let 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 fun1 = Thm.global_cterm_of thy (strip_args 2 t1) + val args1 = map (Thm.global_cterm_of thy) (get_args 2 t1) + val fun2 = Thm.global_cterm_of thy (strip_args 2 t2) + val args2 = map (Thm.global_cterm_of thy) (get_args 1 t2) val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI} 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)) + val subst = map (apfst (Thm.global_cterm_of thy o Var)) (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) in Drule.cterm_instantiate subst relcomppI end @@ -129,10 +129,10 @@ fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT) val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm val typ = Thm.typ_of_cterm rel - val POS_const = Thm.cterm_of thy (mk_POS typ) - val var = Thm.cterm_of thy (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) + val POS_const = Thm.global_cterm_of thy (mk_POS typ) + val var = Thm.global_cterm_of thy (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) val goal = - Thm.apply (Thm.cterm_of thy HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var) + Thm.apply (Thm.global_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 @@ -216,7 +216,7 @@ 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 (Thm.cterm_of thy (Free (hd new_var_names, T))) + val refl_thm = Thm.reflexive (Thm.global_cterm_of thy (Free (hd new_var_names, T))) in Thm.combination def refl_thm |> singleton (Proof_Context.export ctxt' ctxt) @@ -301,7 +301,7 @@ SOME mono_eq_thm => let val rep_abs_eq = mono_eq_thm RS rep_abs_thm - val rep = (Thm.cterm_of thy o quot_thm_rep) quot_thm + val rep = (Thm.global_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 @@ -380,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 |> Thm.cterm_of thy |> Drule.mk_term + fun mk_type typ = typ |> Logic.mk_type |> Thm.global_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 @@ -497,7 +497,7 @@ in fun mk_readable_rsp_thm_eq tm lthy = let - val ctm = Proof_Context.cterm_of lthy tm + val ctm = Thm.cterm_of lthy tm fun assms_rewr_conv tactic rule ct = let @@ -622,7 +622,7 @@ val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (Raw_Simplifier.rewrite lthy false (get_cr_pcr_eqs lthy))) val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs) - |> Proof_Context.cterm_of lthy + |> Thm.cterm_of lthy |> cr_to_pcr_conv |> `Thm.concl_of |>> Logic.dest_equals diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Lifting/lifting_setup.ML Fri Mar 06 15:58:56 2015 +0100 @@ -109,7 +109,7 @@ val (fresh_var, ctxt) = yield_singleton Variable.invent_types sort ctxt val thy = Proof_Context.theory_of ctxt in - ((Thm.cterm_of thy var, Thm.cterm_of thy (HOLogic.eq_const (TFree fresh_var))), ctxt) + ((Thm.global_cterm_of thy var, Thm.global_cterm_of thy (HOLogic.eq_const (TFree fresh_var))), ctxt) end val orig_lthy = lthy @@ -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 (Thm.cterm_of thy fixed_goal) + val init_goal = Goal.init (Thm.global_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 (Thm.cterm_of thy fixed_goal) + val init_goal = Goal.init (Thm.global_cterm_of thy fixed_goal) val rules = Transfer.get_transfer_raw ctxt val rules = constraint :: OO_rules @ rules val tac = @@ -382,7 +382,7 @@ |> 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 (Thm.prop_of id_transfer) [])) val thy = Proof_Context.theory_of lthy - val inst = [(Thm.cterm_of thy var, Thm.cterm_of thy parametrized_relator)] + val inst = [(Thm.global_cterm_of thy var, Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Lifting/lifting_term.ML --- a/src/HOL/Tools/Lifting/lifting_term.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Lifting/lifting_term.ML Fri Mar 06 15:58:56 2015 +0100 @@ -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)) = - (Thm.ctyp_of thy (TVar (x, S)), Thm.ctyp_of thy ty) + (Thm.global_ctyp_of thy (TVar (x, S)), Thm.global_ctyp_of thy ty) val ty_inst = Vartab.fold (cons o (prep_ty thy)) match_env [] in Thm.instantiate (ty_inst, []) quot_thm @@ -373,7 +373,7 @@ let val thy = Proof_Context.theory_of ctxt val instantiated_id_quot_thm = - instantiate' [SOME (Thm.ctyp_of thy ty)] [] @{thm identity_quotient} + instantiate' [SOME (Thm.global_ctyp_of thy ty)] [] @{thm identity_quotient} in (instantiated_id_quot_thm, (table, ctxt)) end @@ -390,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 (Thm.cterm_of thy Q_t) + val Q_thm = Thm.assume (Thm.global_cterm_of thy Q_t) val table' = (ty, Q_thm)::table in (Q_thm, (table', ctxt')) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Meson/meson.ML --- a/src/HOL/Tools/Meson/meson.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Meson/meson.ML Fri Mar 06 15:58:56 2015 +0100 @@ -107,7 +107,7 @@ val tenv = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty) |> snd - val ct_pairs = map (apply2 (Thm.cterm_of thy) o term_pair_of) (Vartab.dest tenv) + val ct_pairs = map (apply2 (Thm.global_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])) @@ -174,7 +174,7 @@ case (Thm.concl_of st, Thm.prop_of th) of (@{const Trueprop} $ (Var _ $ (c as Free _)), @{const Trueprop} $ _) => let - val cc = Thm.cterm_of (Thm.theory_of_thm th) c + val cc = Thm.global_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 @@ -355,7 +355,7 @@ val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (Thm.concl_of th))] ctxt; val spec' = spec - |> Thm.instantiate ([], [(spec_var, Proof_Context.cterm_of ctxt' (Free (x, spec_varT)))]); + |> Thm.instantiate ([], [(spec_var, Thm.cterm_of ctxt' (Free (x, spec_varT)))]); in (th RS spec', ctxt') end end; @@ -629,7 +629,7 @@ $ (t as _ $ _) $ (u as _ $ _))) => (case get_F_pattern T t u of SOME p => - let val inst = [apply2 (Thm.cterm_of thy) (F_ext_cong_neq, p)] in + let val inst = [apply2 (Thm.global_cterm_of thy) (F_ext_cong_neq, p)] in th RS cterm_instantiate inst ext_cong_neq end | NONE => th) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Meson/meson_clausify.ML Fri Mar 06 15:58:56 2015 +0100 @@ -35,8 +35,8 @@ (**** Transformation of Elimination Rules into First-Order Formulas****) -val cfalse = Thm.cterm_of @{theory HOL} @{term False}; -val ctp_false = Thm.cterm_of @{theory HOL} (HOLogic.mk_Trueprop @{term False}); +val cfalse = Thm.global_cterm_of @{theory HOL} @{term False}; +val ctp_false = Thm.global_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 @@ -45,9 +45,9 @@ fun transform_elim_theorem 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 + Thm.instantiate ([], [(Thm.global_cterm_of @{theory HOL} v, cfalse)]) th | v as Var(_, @{typ prop}) => - Thm.instantiate ([], [(Thm.cterm_of @{theory HOL} v, ctp_false)]) th + Thm.instantiate ([], [(Thm.global_cterm_of @{theory HOL} v, ctp_false)]) th | _ => th) @@ -94,9 +94,9 @@ | is_quasi_lambda_free (Abs _) = false | is_quasi_lambda_free _ = true -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})); +val [f_B,g_B] = map (Thm.global_cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_B})); +val [g_C,f_C] = map (Thm.global_cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_C})); +val [f_S,g_S] = map (Thm.global_cterm_of @{theory}) (Misc_Legacy.term_vars (Thm.prop_of @{thm abs_S})); (* FIXME: Requires more use of cterm constructors. *) fun abstract ct = @@ -104,10 +104,10 @@ 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_cterm ct - val cxT = Thm.ctyp_of thy xT - val cbodyT = Thm.ctyp_of thy bodyT + val cxT = Thm.global_ctyp_of thy xT + val cbodyT = Thm.global_ctyp_of thy bodyT fun makeK () = - instantiate' [SOME cxT, SOME cbodyT] [SOME (Thm.cterm_of thy body)] + instantiate' [SOME cxT, SOME cbodyT] [SOME (Thm.global_cterm_of thy body)] @{thm abs_K} in case body of @@ -118,17 +118,17 @@ | rator$rand => if Term.is_dependent rator then (*C or S*) if Term.is_dependent rand then (*S*) - let val crator = Thm.cterm_of thy (Abs(x,xT,rator)) - val crand = Thm.cterm_of thy (Abs(x,xT,rand)) + let val crator = Thm.global_cterm_of thy (Abs(x,xT,rator)) + val crand = Thm.global_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 (Thm.cprop_of abs_S') in Thm.transitive abs_S' (Conv.binop_conv abstract rhs) end else (*C*) - let val crator = Thm.cterm_of thy (Abs(x,xT,rator)) + let val crator = Thm.global_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} + cterm_instantiate [(f_C,crator),(g_C,Thm.global_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) @@ -136,8 +136,8 @@ else if Term.is_dependent rand then (*B or eta*) if rand = Bound 0 then Thm.eta_conversion ct else (*B*) - let val crand = Thm.cterm_of thy (Abs(x,xT,rand)) - val crator = Thm.cterm_of thy rator + let val crand = Thm.global_cterm_of thy (Abs(x,xT,rand)) + val crator = Thm.global_cterm_of thy rator val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B} val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end @@ -179,7 +179,7 @@ TrueI) (*cterms are used throughout for efficiency*) -val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop; +val cTrueprop = Thm.global_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.*) @@ -194,7 +194,7 @@ fun old_skolem_theorem_of_def ctxt rhs0 = let val thy = Proof_Context.theory_of ctxt - val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of thy + val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.global_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 |>> Thm.term_of @@ -202,7 +202,7 @@ case hilbert of Const (_, Type (@{type_name fun}, [_, T])) => T | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) - val cex = Thm.cterm_of thy (HOLogic.exists_const T) + val cex = Thm.global_cterm_of thy (HOLogic.exists_const T) val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) val conc = Drule.list_comb (rhs, frees) @@ -334,7 +334,7 @@ |> (if no_choice then Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of else - Thm.cterm_of thy) + Thm.global_cterm_of thy) |> zap true val fixes = [] |> Term.add_free_names (Thm.prop_of zapped_th) @@ -350,7 +350,7 @@ let val (fully_skolemized_ct, ctxt) = Variable.import_terms true [fully_skolemized_t] ctxt - |>> the_single |>> Thm.cterm_of thy + |>> the_single |>> Thm.global_cterm_of thy in (SOME (discharger_th, fully_skolemized_ct), (Thm.assume fully_skolemized_ct, ctxt)) @@ -377,7 +377,7 @@ th ctxt val (cnf_ths, ctxt) = clausify nnf_th fun intr_imp ct th = - Thm.instantiate ([], map (apply2 (Thm.cterm_of thy)) + Thm.instantiate ([], map (apply2 (Thm.global_cterm_of thy)) [(Var (("i", 0), @{typ nat}), HOLogic.mk_nat ax_no)]) (zero_var_indexes @{thm skolem_COMBK_D}) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Metis/metis_reconstruct.ML --- a/src/HOL/Tools/Metis/metis_reconstruct.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Metis/metis_reconstruct.ML Fri Mar 06 15:58:56 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 = Thm.cterm_of thy o map_types (Logic.incr_tvar idx) +fun cterm_incr_types thy idx = Thm.global_cterm_of thy o map_types (Logic.incr_tvar idx) (* INFERENCE RULE: AXIOM *) @@ -104,7 +104,7 @@ let val th = EXCLUDED_MIDDLE val [vx] = Term.add_vars (Thm.prop_of th) [] - val substs = [(Thm.cterm_of thy (Var vx), Thm.cterm_of thy i_atom)] + val substs = [(Thm.global_cterm_of thy (Var vx), Thm.global_cterm_of thy i_atom)] in cterm_instantiate substs th end @@ -131,7 +131,7 @@ (* We call "polish_hol_terms" below. *) val t = hol_term_of_metis ctxt type_enc sym_tab y in - SOME (Thm.cterm_of thy (Var v), t) + SOME (Thm.global_cterm_of thy (Var v), t) end handle Option.Option => (trace_msg ctxt (fn () => @@ -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 (Thm.ctyp_of thy) (TVar ((a, i), s), TVar ((a, i + inc), s)) + fun inc_tvar ((a, i), s) = apply2 (Thm.global_ctyp_of thy) (TVar ((a, i), s), TVar ((a, i + inc), s)) in Thm.instantiate (map inc_tvar tvs, []) th end @@ -213,7 +213,7 @@ |> rpair (Envir.empty ~1) |-> fold (Pattern.unify (Context.Proof ctxt)) |> Envir.type_env |> Vartab.dest - |> map (fn (x, (S, T)) => apply2 (Thm.ctyp_of thy) (TVar (x, S), T)) + |> map (fn (x, (S, T)) => apply2 (Thm.global_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 @@ -303,7 +303,7 @@ val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp} -val refl_x = Thm.cterm_of @{theory} (Var (hd (Term.add_vars (Thm.prop_of REFL_THM) []))); +val refl_x = Thm.global_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,7 +374,7 @@ 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 (Thm.cterm_of thy)) + val eq_terms = map (apply2 (Thm.global_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' @@ -402,8 +402,8 @@ val thy = Thm.theory_of_thm th val (tyenv, tenv) = fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty) - fun mkT (v, (S, T)) = apply2 (Thm.ctyp_of thy) (TVar (v, S), T) - fun mk (v, (T, t)) = apply2 (Thm.cterm_of thy) (Var (v, Envir.subst_type tyenv T), t) + fun mkT (v, (S, T)) = apply2 (Thm.global_ctyp_of thy) (TVar (v, S), T) + fun mk (v, (T, t)) = apply2 (Thm.global_cterm_of thy) (Var (v, Envir.subst_type tyenv T), t) val instsT = Vartab.fold (cons o mkT) tyenv [] val insts = Vartab.fold (cons o mk) tenv [] @@ -520,7 +520,7 @@ | _ => I) val t_inst = - [] |> try (unify_terms (prem, concl) #> map (apply2 (Thm.cterm_of thy))) + [] |> try (unify_terms (prem, concl) #> map (apply2 (Thm.global_cterm_of thy))) |> the_default [] (* FIXME *) in cterm_instantiate t_inst th @@ -574,9 +574,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 (Thm.ctyp_of thy) (TVar (x, S), T))) + Vartab.fold (fn (x, (S, T)) => cons (apply2 (Thm.global_ctyp_of thy) (TVar (x, S), T))) tyenv [] - val t_inst = [apply2 (Thm.cterm_of thy o Envir.norm_term env) (Var var, t')] + val t_inst = [apply2 (Thm.global_cterm_of thy o Envir.norm_term env) (Var var, t')] in Drule.instantiate_normalize (ty_inst, t_inst) th end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Metis/metis_tactic.ML --- a/src/HOL/Tools/Metis/metis_tactic.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Metis/metis_tactic.ML Fri Mar 06 15:58:56 2015 +0100 @@ -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 = Thm.cterm_of thy t + val ct = Thm.global_cterm_of thy t val cT = Thm.ctyp_of_cterm 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 |> Thm.cterm_of thy |> Thm.trivial + |> HOLogic.mk_Trueprop |> Thm.global_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 = Thm.cterm_of thy (HOLogic.mk_Trueprop t) + val ct = Thm.global_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] @@ -91,7 +91,7 @@ |> (fn th => Meson.first_order_resolve th @{thm Metis.eq_lambdaI}) | v :: _ => Abs (Name.uu, fastype_of v, abstract_over (v, Thm.term_of ct)) $ v - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy |> Conv.comb_conv (conv true ctxt)) else Conv.abs_conv (conv false o snd) ctxt ct @@ -101,7 +101,7 @@ (* We replace the equation's left-hand side with a beta-equivalent term so that "Thm.equal_elim" works below. *) val t0 $ _ $ t2 = Thm.prop_of eq_th - val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.cterm_of thy + val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Mar 06 15:58:56 2015 +0100 @@ -2149,7 +2149,7 @@ SOME wf => wf | NONE => let - val goal = prop |> Thm.cterm_of thy |> Goal.init + val goal = prop |> Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Nitpick/nitpick_model.ML --- a/src/HOL/Tools/Nitpick/nitpick_model.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Fri Mar 06 15:58:56 2015 +0100 @@ -1062,7 +1062,7 @@ |> writeln else () - val goal = prop |> Thm.cterm_of thy |> Goal.init + val goal = prop |> Thm.global_cterm_of thy |> Goal.init in (goal |> SINGLE (DETERM_TIMEOUT auto_timeout (auto_tac ctxt)) |> the |> Goal.finish ctxt; true) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Old_Datatype/old_datatype.ML --- a/src/HOL/Tools/Old_Datatype/old_datatype.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML Fri Mar 06 15:58:56 2015 +0100 @@ -247,9 +247,9 @@ (thy, defs, eqns, rep_congs, dist_lemmas) = let 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 rep_const = Thm.global_cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT)); + val cong' = cterm_instantiate [(Thm.global_cterm_of thy cong_f, rep_const)] arg_cong; + val dist = cterm_instantiate [(Thm.global_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 (Thm.ctyp_of thy5 T)] [] fun_cong)) branchTs; + map (fn T => make_elim (Drule.instantiate' [SOME (Thm.global_ctyp_of thy5 T)] [] fun_cong)) branchTs; fun prove_iso_thms ds (inj_thms, elem_thms) = let @@ -626,7 +626,7 @@ 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 (Thm.cterm_of thy6) Ps ~~ map (Thm.cterm_of thy6) frees) indrule_lemma; + cterm_instantiate (map (Thm.global_cterm_of thy6) Ps ~~ map (Thm.global_cterm_of thy6) frees) indrule_lemma; val dt_induct_prop = Old_Datatype_Prop.make_ind descr; val dt_induct = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Old_Datatype/old_datatype_aux.ML --- a/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Old_Datatype/old_datatype_aux.ML Fri Mar 06 15:58:56 2015 +0100 @@ -149,7 +149,7 @@ (case abstr (getP u) of NONE => NONE | SOME u' => - SOME (apply2 (Thm.cterm_of thy) (t |> getP |> snd |> head_of, u')))) (ts ~~ ts'); + SOME (apply2 (Thm.global_cterm_of thy) (t |> getP |> snd |> head_of, u')))) (ts ~~ ts'); val indrule' = cterm_instantiate insts indrule; in resolve0_tac [indrule'] i end); @@ -166,7 +166,7 @@ val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem')); val exhaustion' = cterm_instantiate - [apply2 (Proof_Context.cterm_of ctxt) + [apply2 (Thm.cterm_of ctxt) (head_of lhs, 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); diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Old_Datatype/old_rep_datatype.ML --- a/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Old_Datatype/old_rep_datatype.ML Fri Mar 06 15:58:56 2015 +0100 @@ -46,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 insts' = map (Thm.cterm_of thy) induct_Ps ~~ map (Thm.cterm_of thy) insts; + val insts' = map (Thm.global_cterm_of thy) induct_Ps ~~ map (Thm.global_cterm_of thy) insts; val induct' = refl RS (nth (Old_Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i @@ -208,7 +208,7 @@ ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts); val induct' = induct |> cterm_instantiate - (map (Thm.cterm_of thy1) induct_Ps ~~ map (Thm.cterm_of thy1) insts); + (map (Thm.global_cterm_of thy1) induct_Ps ~~ map (Thm.global_cterm_of thy1) insts); in Old_Datatype_Aux.split_conj_thm (Goal.prove_sorry_global thy1 [] [] (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj rec_unique_ts)) @@ -381,7 +381,7 @@ let val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (Thm.prems_of exhaustion))); val exhaustion' = exhaustion - |> cterm_instantiate [apply2 (Thm.cterm_of thy) (lhs, Free ("x", T))]; + |> cterm_instantiate [apply2 (Thm.global_cterm_of thy) (lhs, Free ("x", T))]; fun tac ctxt = EVERY [resolve_tac ctxt [exhaustion'] 1, ALLGOALS (asm_simp_tac @@ -453,7 +453,7 @@ val nchotomy' = nchotomy RS spec; val [v] = Term.add_vars (Thm.concl_of nchotomy') []; val nchotomy'' = - cterm_instantiate [apply2 (Thm.cterm_of thy) (Var v, Ma)] nchotomy'; + cterm_instantiate [apply2 (Thm.global_cterm_of thy) (Var v, Ma)] nchotomy'; in Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) (fn {context = ctxt, prems, ...} => diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Predicate_Compile/core_data.ML --- a/src/HOL/Tools/Predicate_Compile/core_data.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Predicate_Compile/core_data.ML Fri Mar 06 15:58:56 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 (Thm.cterm_of thy) o term_pair_of) + |> snd |> Vartab.dest |> map (apply2 (Thm.global_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)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Fri Mar 06 15:58:56 2015 +0100 @@ -877,7 +877,7 @@ 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, Thm.cterm_of thy (Pattern.rewrite_term thy pats' [] (Thm.term_of ct2))) + (ct1, Thm.global_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''] diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Mar 06 15:58:56 2015 +0100 @@ -158,14 +158,14 @@ in (ps, Logic.list_implies (prems, HOLogic.mk_Trueprop lhs)) end - val x = (Thm.cterm_of thy o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o + val x = (Thm.global_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, Thm.cterm_of thy (Free y))] @{thm exI}) 1) ps)) + rtac (Drule.cterm_instantiate [(x, Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Qelim/cooper.ML Fri Mar 06 15:58:56 2015 +0100 @@ -694,7 +694,7 @@ val (_, oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (@{binding cooper}, (fn (ctxt, t) => - (Proof_Context.cterm_of ctxt o Logic.mk_equals o apply2 HOLogic.mk_Trueprop) + (Thm.cterm_of ctxt o Logic.mk_equals o apply2 HOLogic.mk_Trueprop) (t, procedure t))))); val comp_ss = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Quickcheck/narrowing_generators.ML --- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML Fri Mar 06 15:58:56 2015 +0100 @@ -75,9 +75,9 @@ (map mk_partial_term_of (frees ~~ tys)) (@{term "Code_Evaluation.Const"} $ HOLogic.mk_literal c $ HOLogic.mk_typerep (tys ---> ty)) val insts = - map (SOME o Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) + map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) [Free ("ty", Term.itselfT ty), narrowing_term, rhs] - val cty = Thm.ctyp_of thy ty; + val cty = Thm.global_ctyp_of thy ty; in @{thm partial_term_of_anything} |> Drule.instantiate' [SOME cty] insts @@ -94,12 +94,12 @@ (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs; val const = Axclass.param_of_inst thy (@{const_name partial_term_of}, tyco); val var_insts = - map (SOME o Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) + map (SOME o Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) [Free ("ty", Term.itselfT ty), @{term "Quickcheck_Narrowing.Narrowing_variable p tt"}, @{term "Code_Evaluation.Free (STR ''_'')"} $ HOLogic.mk_typerep ty]; val var_eq = @{thm partial_term_of_anything} - |> Drule.instantiate' [SOME (Thm.ctyp_of thy ty)] var_insts + |> Drule.instantiate' [SOME (Thm.global_ctyp_of thy ty)] var_insts |> Thm.varifyT_global val eqs = var_eq :: map_index (mk_partial_term_of_eq thy ty) cs; in diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Quickcheck/random_generators.ML --- a/src/HOL/Tools/Quickcheck/random_generators.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Quickcheck/random_generators.ML Fri Mar 06 15:58:56 2015 +0100 @@ -88,7 +88,7 @@ val ((t_random_aux as Free (random_aux, T)) $ (t_k as Free (v, _)), proto_t_rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq; val Type (_, [_, iT]) = T; - val icT = Thm.ctyp_of thy iT; + val icT = Thm.global_ctyp_of thy iT; val inst = Thm.instantiate_cterm ([(aT, icT)], []); fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t); val t_rhs = lambda t_k proto_t_rhs; @@ -103,7 +103,7 @@ val rule = @{thm random_aux_rec} |> Drule.instantiate_normalize ([(aT, icT)], - [(cT_random_aux, Thm.cterm_of thy t_random_aux), (cT_rhs, Thm.cterm_of thy t_rhs)]); + [(cT_random_aux, Thm.global_cterm_of thy t_random_aux), (cT_rhs, Thm.global_cterm_of thy t_rhs)]); fun tac ctxt = ALLGOALS (rtac rule) THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Quotient/quotient_def.ML --- a/src/HOL/Tools/Quotient/quotient_def.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Quotient/quotient_def.ML Fri Mar 06 15:58:56 2015 +0100 @@ -91,7 +91,7 @@ fun mk_readable_rsp_thm_eq tm lthy = let - val ctm = Proof_Context.cterm_of lthy tm + val ctm = Thm.cterm_of lthy tm fun norm_fun_eq ctm = let diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Quotient/quotient_tacs.ML --- a/src/HOL/Tools/Quotient/quotient_tacs.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Fri Mar 06 15:58:56 2015 +0100 @@ -73,10 +73,10 @@ fun prep_trm thy (x, (T, t)) = - (Thm.cterm_of thy (Var (x, T)), Thm.cterm_of thy t) + (Thm.global_cterm_of thy (Var (x, T)), Thm.global_cterm_of thy t) fun prep_ty thy (x, (S, ty)) = - (Thm.ctyp_of thy (TVar (x, S)), Thm.ctyp_of thy ty) + (Thm.global_ctyp_of thy (TVar (x, S)), Thm.global_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 Thm.ctyp_of thy) [domain_type (fastype_of R2)] - val trm_inst = map (SOME o Thm.cterm_of thy) [R2, R1] + val ty_inst = map (SOME o Thm.global_ctyp_of thy) [domain_type (fastype_of R2)] + val trm_inst = map (SOME o Thm.global_cterm_of thy) [R2, R1] in (case try (Drule.instantiate' ty_inst trm_inst) ball_bex_thm of NONE => NONE @@ -199,10 +199,10 @@ let val fx = fnctn x; val thy = Proof_Context.theory_of ctxt; - 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 cx = Thm.global_cterm_of thy x; + val cfx = Thm.global_cterm_of thy fx; + val cxt = Thm.global_ctyp_of thy (fastype_of x); + val cfxt = Thm.global_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 @@ -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 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 ty_inst = map (SOME o Thm.global_ctyp_of thy) [ty_x, ty_b, ty_f] + val t_inst = map (SOME o Thm.global_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 (Thm.cterm_of thy) R of (* There can be loose bounds in R *) + case try (Thm.global_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 (Thm.ctyp_of thy ty)] - [SOME (Thm.cterm_of thy R)]) @{thm equals_rsp} of + case try (Drule.instantiate' [SOME (Thm.global_ctyp_of thy ty)] + [SOME (Thm.global_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 Thm.ctyp_of thy) [ty_a, ty_b]; + val ty_inst = map (SOME o Thm.global_ctyp_of thy) [ty_a, ty_b]; in - case try (map (SOME o Thm.cterm_of thy)) [rel, abs, rep] of + case try (map (SOME o Thm.global_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 @@ -478,8 +478,8 @@ 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 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 tyinst = map (SOME o Thm.global_ctyp_of thy) [ty_a, ty_b, ty_c, ty_d] + val tinst = [NONE, NONE, SOME (Thm.global_cterm_of thy r1), NONE, SOME (Thm.global_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 @@ -488,7 +488,7 @@ 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 + Drule.instantiate_normalize ([], [(Thm.global_cterm_of thy insp, Thm.global_cterm_of thy inst)]) thm3 in Conv.rewr_conv thm4 ctrm end @@ -557,7 +557,7 @@ let val thy = Proof_Context.theory_of ctxt val vrs = Term.add_frees concl [] - val cvrs = map (Thm.cterm_of thy o Free) vrs + val cvrs = map (Thm.global_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 @@ -617,10 +617,10 @@ handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Drule.instantiate' [] - [SOME (Thm.cterm_of thy rtrm'), - SOME (Thm.cterm_of thy reg_goal), + [SOME (Thm.global_cterm_of thy rtrm'), + SOME (Thm.global_cterm_of thy reg_goal), NONE, - SOME (Thm.cterm_of thy inj_goal)] procedure_thm + SOME (Thm.global_cterm_of thy inj_goal)] procedure_thm end @@ -677,9 +677,9 @@ handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Drule.instantiate' [] - [SOME (Thm.cterm_of thy reg_goal), + [SOME (Thm.global_cterm_of thy reg_goal), NONE, - SOME (Thm.cterm_of thy inj_goal)] partiality_procedure_thm + SOME (Thm.global_cterm_of thy inj_goal)] partiality_procedure_thm end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/smt_normalize.ML Fri Mar 06 15:58:56 2015 +0100 @@ -31,7 +31,7 @@ (** instantiate elimination rules **) local - val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.cterm_of @{theory} @{const False}) + val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.global_cterm_of @{theory} @{const False}) fun inst f ct thm = let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) @@ -189,7 +189,7 @@ fun mk_pat ct = Thm.apply (SMT_Util.instT' ct pat) ct fun mk_clist T = - apply2 (Thm.cterm_of @{theory}) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T) + apply2 (Thm.global_cterm_of @{theory}) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T) fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil val mk_pat_list = mk_list (mk_clist @{typ pattern}) val mk_mpat_list = mk_list (mk_clist @{typ "pattern symb_list"}) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/smt_real.ML --- a/src/HOL/Tools/SMT/smt_real.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/smt_real.ML Fri Mar 06 15:58:56 2015 +0100 @@ -64,14 +64,14 @@ fun mk_nary _ cu [] = cu | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts) - val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (real)}) - val add = Thm.cterm_of @{theory} @{const plus (real)} + val mk_uminus = Thm.apply (Thm.global_cterm_of @{theory} @{const uminus (real)}) + val add = Thm.global_cterm_of @{theory} @{const plus (real)} val real0 = Numeral.mk_cnumber @{ctyp real} 0 - val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (real)}) - val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (real)}) - val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const divide (real)}) - val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (real)}) - val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (real)}) + val mk_sub = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const minus (real)}) + val mk_mul = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const times (real)}) + val mk_div = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const divide (real)}) + val mk_lt = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less (real)}) + val mk_le = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less_eq (real)}) fun z3_mk_builtin_fun (Z3_Interface.Sym ("-", _)) [ct] = SOME (mk_uminus ct) | z3_mk_builtin_fun (Z3_Interface.Sym ("+", _)) cts = SOME (mk_nary add real0 cts) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/smt_solver.ML --- a/src/HOL/Tools/SMT/smt_solver.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/smt_solver.ML Fri Mar 06 15:58:56 2015 +0100 @@ -207,7 +207,7 @@ else oracle () | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat))) - val cfalse = Thm.cterm_of @{theory} @{prop False} + val cfalse = Thm.global_cterm_of @{theory} @{prop False} in fun add_solver ({name, class, avail, command, options, smt_options, default_max_relevant, outcome, diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/smt_util.ML --- a/src/HOL/Tools/SMT/smt_util.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/smt_util.ML Fri Mar 06 15:58:56 2015 +0100 @@ -165,7 +165,7 @@ (* patterns and instantiations *) fun mk_const_pat thy name destT = - let val cpat = Thm.cterm_of thy (Const (name, Sign.the_const_type thy name)) + let val cpat = Thm.global_cterm_of thy (Const (name, Sign.the_const_type thy name)) in (destT (Thm.ctyp_of_cterm cpat), cpat) end val destT1 = hd o Thm.dest_ctyp @@ -194,7 +194,7 @@ val dest_all_cbinders = repeat_yield (try o dest_cbinder) -val mk_cprop = Thm.apply (Thm.cterm_of @{theory} @{const Trueprop}) +val mk_cprop = Thm.apply (Thm.global_cterm_of @{theory} @{const Trueprop}) fun dest_cprop ct = (case Thm.term_of ct of diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/z3_interface.ML --- a/src/HOL/Tools/SMT/z3_interface.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/z3_interface.ML Fri Mar 06 15:58:56 2015 +0100 @@ -112,13 +112,13 @@ | mk_builtin_num ctxt i T = chained_mk_builtin_num ctxt (get_mk_builtins ctxt) i T -val mk_true = Thm.cterm_of @{theory} (@{const Not} $ @{const False}) -val mk_false = Thm.cterm_of @{theory} @{const False} -val mk_not = Thm.apply (Thm.cterm_of @{theory} @{const Not}) -val mk_implies = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.implies}) -val mk_iff = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.eq (bool)}) -val conj = Thm.cterm_of @{theory} @{const HOL.conj} -val disj = Thm.cterm_of @{theory} @{const HOL.disj} +val mk_true = Thm.global_cterm_of @{theory} (@{const Not} $ @{const False}) +val mk_false = Thm.global_cterm_of @{theory} @{const False} +val mk_not = Thm.apply (Thm.global_cterm_of @{theory} @{const Not}) +val mk_implies = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const HOL.implies}) +val mk_iff = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const HOL.eq (bool)}) +val conj = Thm.global_cterm_of @{theory} @{const HOL.conj} +val disj = Thm.global_cterm_of @{theory} @{const HOL.disj} fun mk_nary _ cu [] = cu | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts) @@ -139,15 +139,15 @@ let val cTs = Thm.dest_ctyp (Thm.ctyp_of_cterm array) in Thm.apply (Thm.mk_binop (SMT_Util.instTs cTs update) array index) value end -val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (int)}) -val add = Thm.cterm_of @{theory} @{const plus (int)} +val mk_uminus = Thm.apply (Thm.global_cterm_of @{theory} @{const uminus (int)}) +val add = Thm.global_cterm_of @{theory} @{const plus (int)} val int0 = Numeral.mk_cnumber @{ctyp int} 0 -val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (int)}) -val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (int)}) -val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3div}) -val mk_mod = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3mod}) -val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (int)}) -val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (int)}) +val mk_sub = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const minus (int)}) +val mk_mul = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const times (int)}) +val mk_div = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const z3div}) +val mk_mod = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const z3mod}) +val mk_lt = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less (int)}) +val mk_le = Thm.mk_binop (Thm.global_cterm_of @{theory} @{const less_eq (int)}) fun mk_builtin_fun ctxt sym cts = (case (sym, cts) of diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/z3_replay.ML --- a/src/HOL/Tools/SMT/z3_replay.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/z3_replay.ML Fri Mar 06 15:58:56 2015 +0100 @@ -23,7 +23,7 @@ val maxidx = Thm.maxidx_of thm + 1 val vs = params_of (Thm.prop_of thm) val vars = map_index (fn (i, (n, T)) => Var ((n, i + maxidx), T)) vs - in Drule.forall_elim_list (map (Proof_Context.cterm_of ctxt) vars) thm end + in Drule.forall_elim_list (map (Thm.cterm_of ctxt) vars) thm end fun add_paramTs names t = fold2 (fn n => fn (_, T) => AList.update (op =) (n, T)) names (params_of t) @@ -31,7 +31,7 @@ fun new_fixes ctxt nTs = let val (ns, ctxt') = Variable.variant_fixes (replicate (length nTs) "") ctxt - fun mk (n, T) n' = (n, Proof_Context.cterm_of ctxt' (Free (n', T))) + fun mk (n, T) n' = (n, Thm.cterm_of ctxt' (Free (n', T))) in (ctxt', Symtab.make (map2 mk nTs ns)) end fun forall_elim_term ct (Const (@{const_name Pure.all}, _) $ (a as Abs _)) = @@ -60,7 +60,7 @@ if Z3_Proof.is_assumption rule then (case Inttab.lookup assumed id of SOME (_, thm) => thm - | NONE => Thm.assume (Proof_Context.cterm_of ctxt concl)) + | NONE => Thm.assume (Thm.cterm_of ctxt concl)) else under_fixes (Z3_Replay_Methods.method_for rule) ctxt (if is_fix_step then (map snd nthms, []) else ([], nthms)) fixes concl @@ -123,7 +123,7 @@ (cx as ((iidths, thms), (ctxt, ptab))) = if Z3_Proof.is_assumption rule andalso rule <> Z3_Proof.Hypothesis then let - val ct = Proof_Context.cterm_of ctxt concl + val ct = Thm.cterm_of ctxt concl val thm1 = Thm.trivial ct |> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt)) val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1 in diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/SMT/z3_replay_methods.ML --- a/src/HOL/Tools/SMT/z3_replay_methods.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/SMT/z3_replay_methods.ML Fri Mar 06 15:58:56 2015 +0100 @@ -88,7 +88,7 @@ fun dest_thm thm = dest_prop (Thm.concl_of thm) -fun certify_prop ctxt t = Proof_Context.cterm_of ctxt (as_prop t) +fun certify_prop ctxt t = Thm.cterm_of ctxt (as_prop t) fun try_provers ctxt rule [] thms t = replay_rule_error ctxt rule thms t | try_provers ctxt rule ((name, prover) :: named_provers) thms t = @@ -108,12 +108,12 @@ fun match_instantiateT ctxt t thm = if Term.exists_type (Term.exists_subtype Term.is_TVar) (dest_thm thm) then - Thm.instantiate (gen_certify_inst fst TVar (Proof_Context.ctyp_of ctxt) ctxt thm t, []) thm + Thm.instantiate (gen_certify_inst fst TVar (Thm.ctyp_of ctxt) ctxt thm t, []) thm else thm fun match_instantiate ctxt t thm = let val thm' = match_instantiateT ctxt t thm in - Thm.instantiate ([], gen_certify_inst snd Var (Proof_Context.cterm_of ctxt) ctxt thm' t) thm' + Thm.instantiate ([], gen_certify_inst snd Var (Thm.cterm_of ctxt) ctxt thm' t) thm' end fun apply_rule ctxt t = @@ -400,7 +400,7 @@ end fun forall_intr ctxt t thm = - let val ct = Proof_Context.cterm_of ctxt t + let val ct = Thm.cterm_of ctxt t in Thm.forall_intr ct thm COMP_INCR @{thm iff_allI} end in diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Fri Mar 06 15:58:56 2015 +0100 @@ -807,7 +807,7 @@ | freeze (Free (s, T)) = Free (s, freezeT T) | freeze t = t -fun goal_of_thm thy = Thm.prop_of #> freeze #> Thm.cterm_of thy #> Goal.init +fun goal_of_thm thy = Thm.prop_of #> freeze #> Thm.global_cterm_of thy #> Goal.init fun run_prover_for_mash ctxt params prover goal_name facts goal = let diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/TFL/casesplit.ML --- a/src/HOL/Tools/TFL/casesplit.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/TFL/casesplit.ML Fri Mar 06 15:58:56 2015 +0100 @@ -37,8 +37,8 @@ val x = Free(vstr,ty) val abst = Abs(vstr, ty, Term.abstract_over (x, gt)); - val ctermify = Thm.cterm_of sgn; - val ctypify = Thm.ctyp_of sgn; + val ctermify = Thm.global_cterm_of sgn; + val ctypify = Thm.global_ctyp_of sgn; val case_thm = case_thm_of_ty sgn ty; val abs_ct = ctermify abst; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/TFL/dcterm.ML --- a/src/HOL/Tools/TFL/dcterm.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/TFL/dcterm.ML Fri Mar 06 15:58:56 2015 +0100 @@ -67,7 +67,7 @@ * Some simple constructor functions. *---------------------------------------------------------------------------*) -val mk_hol_const = Thm.cterm_of @{theory HOL} o Const; +val mk_hol_const = Thm.global_cterm_of @{theory HOL} o Const; fun mk_exists (r as (Bvar, Body)) = let val ty = Thm.typ_of_cterm Bvar @@ -181,7 +181,7 @@ val t = Thm.term_of ctm; in if can HOLogic.dest_Trueprop t then ctm - else Thm.cterm_of thy (HOLogic.mk_Trueprop t) + else Thm.global_cterm_of thy (HOLogic.mk_Trueprop t) end handle TYPE (msg, _, _) => raise ERR "mk_prop" msg | TERM (msg, _) => raise ERR "mk_prop" msg; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/TFL/post.ML --- a/src/HOL/Tools/TFL/post.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/TFL/post.ML Fri Mar 06 15:58:56 2015 +0100 @@ -73,7 +73,7 @@ fun join_assums ctxt th = let val thy = Thm.theory_of_thm th - val tych = Thm.cterm_of thy + val tych = Thm.global_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 *) @@ -175,7 +175,7 @@ (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths); in fun derive_init_eqs ctxt rules eqs = - map (Thm.trivial o Proof_Context.cterm_of ctxt o HOLogic.mk_Trueprop) eqs + map (Thm.trivial o Thm.cterm_of ctxt o HOLogic.mk_Trueprop) eqs |> map_index (fn (i, e) => solve_eq ctxt (e, (get_related_thms i rules), i)) |> flat; end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/TFL/rules.ML --- a/src/HOL/Tools/TFL/rules.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/TFL/rules.ML Fri Mar 06 15:58:56 2015 +0100 @@ -161,7 +161,7 @@ local val thy = Thm.theory_of_thm disjI1 val prop = Thm.prop_of disjI1 val [P,Q] = Misc_Legacy.term_vars prop - val disj1 = Thm.forall_intr (Thm.cterm_of thy Q) disjI1 + val disj1 = Thm.forall_intr (Thm.global_cterm_of thy Q) disjI1 in fun DISJ1 thm tm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj1) handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg; @@ -170,7 +170,7 @@ local val thy = Thm.theory_of_thm disjI2 val prop = Thm.prop_of disjI2 val [P,Q] = Misc_Legacy.term_vars prop - val disj2 = Thm.forall_intr (Thm.cterm_of thy P) disjI2 + val disj2 = Thm.forall_intr (Thm.global_cterm_of thy P) disjI2 in fun DISJ2 tm thm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj2) handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg; @@ -265,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 = Thm.ctyp_of thy (type_of x) - val gspec = Thm.forall_intr (Thm.cterm_of thy x) spec + val cTV = Thm.global_ctyp_of thy (type_of x) + val gspec = Thm.forall_intr (Thm.global_cterm_of thy x) spec in fun SPEC tm thm = let val gspec' = Drule.instantiate_normalize ([(cTV, Thm.ctyp_of_cterm tm)], []) gspec @@ -282,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)) => (Thm.ctyp_of s (TVar (i, S)), Thm.ctyp_of s ty)) + fun cty_theta s = map (fn (i, (S, ty)) => (Thm.global_ctyp_of s (TVar (i, S)), Thm.global_ctyp_of s ty)) fun ctm_theta s = map (fn (i, (_, tm2)) => - let val ctm2 = Thm.cterm_of s tm2 - in (Thm.cterm_of s (Var(i, Thm.typ_of_cterm ctm2)), ctm2) + let val ctm2 = Thm.global_cterm_of s tm2 + in (Thm.global_cterm_of s (Var(i, Thm.typ_of_cterm ctm2)), ctm2) end) fun certify s (ty_theta,tm_theta) = (cty_theta s (Vartab.dest ty_theta), @@ -301,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 (Thm.cterm_of thy prop') + in ALPHA thm (Thm.global_cterm_of thy prop') end end; @@ -310,7 +310,7 @@ fun GEN_ALL thm = let val thy = Thm.theory_of_thm thm val prop = Thm.prop_of thm - val tycheck = Thm.cterm_of thy + val tycheck = Thm.global_cterm_of thy val vlist = map tycheck (Misc_Legacy.add_term_vars (prop, [])) in GENL vlist thm end; @@ -343,7 +343,7 @@ val redex = Dcterm.capply lam fvar val thy = Thm.theory_of_cterm redex val t$u = Thm.term_of redex - val residue = Thm.cterm_of thy (Term.betapply (t, u)) + val residue = Thm.global_cterm_of thy (Term.betapply (t, u)) in GEN fvar (DISCH residue fact) RS (exth RS Thms.choose_thm) handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg @@ -356,8 +356,8 @@ fun EXISTS (template,witness) thm = let val thy = Thm.theory_of_thm thm val prop = Thm.prop_of thm - val P' = Thm.cterm_of thy P - val x' = Thm.cterm_of thy x + val P' = Thm.global_cterm_of thy P + val x' = Thm.global_cterm_of thy x val abstr = #2 (Dcterm.dest_comb template) in thm RS (cterm_instantiate[(P',abstr), (x',witness)] exI) @@ -389,9 +389,9 @@ fun IT_EXISTS blist th = let val thy = Thm.theory_of_thm th - val tych = Thm.cterm_of thy + val tych = Thm.global_cterm_of thy val blist' = map (apply2 Thm.term_of) blist - fun ex v M = Thm.cterm_of thy (USyntax.mk_exists{Bvar=v,Body = M}) + fun ex v M = Thm.global_cterm_of thy (USyntax.mk_exists{Bvar=v,Body = M}) in fold_rev (fn (b as (r1,r2)) => fn thm => @@ -511,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 = Thm.cterm_of thy + val tych = Thm.global_cterm_of thy val ants = Logic.strip_imp_prems (Thm.prop_of thm) val news = get (ants,1,[]) in @@ -663,7 +663,7 @@ val dummy = say (Display.string_of_thm ctxt thm) (* Unquantified eliminate *) fun uq_eliminate (thm,imp,thy) = - let val tych = Thm.cterm_of thy + let val tych = Thm.global_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 @@ -683,7 +683,7 @@ orelse error "assertion failed in CONTEXT_REWRITE_RULE" val imp_body1 = subst_free (ListPair.zip (args, vstrl)) imp_body - val tych = Thm.cterm_of thy + val tych = Thm.global_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 @@ -712,7 +712,7 @@ in if (pbeta_redex Q) (length vlist) then pq_eliminate (thm,thy,vlist,imp_body,Q) else - let val tych = Thm.cterm_of thy + let val tych = Thm.global_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 @@ -763,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:" (Thm.cterm_of thy func) - val dummy = print_cterm ctxt "TC:" (Thm.cterm_of thy (HOLogic.mk_Trueprop TC)) + val dummy = print_cterm ctxt "func:" (Thm.global_cterm_of thy func) + val dummy = print_cterm ctxt "TC:" (Thm.global_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 = Thm.cterm_of thy + else let val cTC = Thm.global_cterm_of thy (HOLogic.mk_Trueprop TC) in case rcontext of [] => SPEC_ALL(ASSUME cTC) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/TFL/thry.ML --- a/src/HOL/Tools/TFL/thry.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/TFL/thry.ML Fri Mar 06 15:58:56 2015 +0100 @@ -48,7 +48,7 @@ *---------------------------------------------------------------------------*) fun typecheck thry t = - Thm.cterm_of thry t + Thm.global_cterm_of thry t handle TYPE (msg, _, _) => raise THRY_ERR "typecheck" msg | TERM (msg, _) => raise THRY_ERR "typecheck" msg; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Transfer/transfer.ML --- a/src/HOL/Tools/Transfer/transfer.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Transfer/transfer.ML Fri Mar 06 15:58:56 2015 +0100 @@ -238,7 +238,7 @@ val prems = map (HOLogic.mk_Trueprop o mk_is_equality) frees val prop1 = mk_prop' (Term.subst_atomic (eq_consts ~~ frees) t) val prop2 = fold Logic.all frees (Logic.list_implies (prems, prop1)) - val cprop = Thm.cterm_of thy prop2 + val cprop = Thm.global_cterm_of thy prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false [is_equality_lemma] cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm in @@ -331,7 +331,7 @@ val t' = subst_terms (fold Termtab.update (Domainp_tms ~~ frees) Termtab.empty) t val prop1 = fold Logic.all frees (Logic.list_implies (prems, mk_prop' t')) val prop2 = Logic.list_rename_params (rev names) prop1 - val cprop = Thm.cterm_of thy prop2 + val cprop = Thm.global_cterm_of thy prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false [Domainp_lemma] cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm; in @@ -436,7 +436,7 @@ let val ([x', y'], ctxt') = Variable.variant_fixes [x, y] ctxt val prop = mk_Rel (rel T U) $ Free (x', T) $ Free (y', U) - val cprop = Thm.cterm_of thy (HOLogic.mk_Trueprop prop) + val cprop = Thm.global_cterm_of thy (HOLogic.mk_Trueprop prop) 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)) @@ -462,13 +462,13 @@ val T = fastype_of t val U = fastype_of u val prop = mk_Rel (rel T U) $ t $ u - val cprop = Thm.cterm_of thy (HOLogic.mk_Trueprop prop) + val cprop = Thm.global_cterm_of thy (HOLogic.mk_Trueprop prop) in (Thm.assume cprop, [cprop]) end val r = mk_Rel (rel (fastype_of t) (fastype_of u)) val goal = HOLogic.mk_Trueprop (r $ t $ u) - val rename = Thm.trivial (Thm.cterm_of thy goal) + val rename = Thm.trivial (Thm.global_cterm_of thy goal) val (thm, hyps) = zip ctxt [] t u in Drule.implies_intr_list hyps (thm RS rename) @@ -578,9 +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, _) = (Thm.ctyp_of thy (TVar ((a, idx), @{sort type})), cbool) + fun tinst (a, _) = (Thm.global_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) + (Thm.global_cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), Thm.global_cterm_of thy t) in thm |> Thm.generalize (tfrees, rnames @ frees) idx @@ -604,9 +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, _) = (Thm.ctyp_of thy (TVar ((a, idx), @{sort type})), cbool) + fun tinst (a, _) = (Thm.global_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) + (Thm.global_cterm_of thy (Var (Name.clean_index (prep a, idx), relT)), Thm.global_cterm_of thy t) in thm |> Thm.generalize (tfrees, rnames @ frees) idx diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/Transfer/transfer_bnf.ML --- a/src/HOL/Tools/Transfer/transfer_bnf.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/Transfer/transfer_bnf.ML Fri Mar 06 15:58:56 2015 +0100 @@ -344,8 +344,8 @@ val predTs = map mk_pred1T As val (preds, lthy) = mk_Frees "P" predTs lthy 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 + val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) As) + val cts = map (SOME o Thm.cterm_of lthy) args 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 = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/choice_specification.ML Fri Mar 06 15:58:56 2015 +0100 @@ -83,7 +83,7 @@ val ctxt = Proof_Context.init_global thy val rew_imps = alt_props |> - map (Object_Logic.atomize ctxt o Thm.cterm_of thy o Syntax.read_prop_global thy o snd) + map (Object_Logic.atomize ctxt o Thm.global_cterm_of thy o Syntax.read_prop_global thy o snd) val props' = rew_imps |> map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of) @@ -136,7 +136,7 @@ let fun inst_all thy v thm = let - val cv = Thm.cterm_of thy v + val cv = Thm.global_cterm_of thy v val cT = Thm.ctyp_of_cterm cv val spec' = instantiate' [SOME cT] [NONE, SOME cv] spec in thm RS spec' end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/cnf.ML --- a/src/HOL/Tools/cnf.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/cnf.ML Fri Mar 06 15:58:56 2015 +0100 @@ -165,7 +165,7 @@ (* ------------------------------------------------------------------------- *) fun inst_thm thy ts thm = - instantiate' [] (map (SOME o Thm.cterm_of thy) ts) thm; + instantiate' [] (map (SOME o Thm.global_cterm_of thy) ts) thm; (* ------------------------------------------------------------------------- *) (* Naive CNF transformation *) @@ -266,7 +266,7 @@ | Abs _ => Conv.abs_conv (conv o snd) ctxt ct | Const _ => Conv.all_conv ct | t => make t RS eq_reflection) - in conv ctxt (Thm.cterm_of thy t) RS meta_eq_to_obj_eq end + in conv ctxt (Thm.global_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)) @@ -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 (Thm.cterm_of thy var) thm2 (* !!v. (x' | y') = body' *) + val thm3 = Thm.forall_intr (Thm.global_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 (Thm.cterm_of thy var) thm2 (* !!v. (x' | y') = body' *) + val thm3 = Thm.forall_intr (Thm.global_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 @@ -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 (Thm.cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *) + val thm3 = Thm.forall_intr (Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/code_evaluation.ML --- a/src/HOL/Tools/code_evaluation.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/code_evaluation.ML Fri Mar 06 15:58:56 2015 +0100 @@ -62,11 +62,11 @@ val t = list_comb (Const (c, tys ---> ty), map Free (Name.invent_names Name.context "a" tys)); val (arg, rhs) = - apply2 (Thm.cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) + apply2 (Thm.global_cterm_of thy o Logic.unvarify_types_global o Logic.varify_global) (t, map_aterms (fn t as Free (_, ty) => HOLogic.mk_term_of ty t | t => t) (HOLogic.reflect_term t)); - val cty = Thm.ctyp_of thy ty; + val cty = Thm.global_ctyp_of thy ty; in @{thm term_of_anything} |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs] @@ -102,11 +102,11 @@ val arg = Var (("x", 0), ty); val rhs = Abs ("y", @{typ term}, HOLogic.reflect_term (Const (abs, ty_rep --> ty) $ Bound 0)) $ (HOLogic.mk_term_of ty_rep (Const (proj, ty --> ty_rep) $ arg)) - |> Thm.cterm_of thy; - val cty = Thm.ctyp_of thy ty; + |> Thm.global_cterm_of thy; + val cty = Thm.global_ctyp_of thy ty; in @{thm term_of_anything} - |> Drule.instantiate' [SOME cty] [SOME (Thm.cterm_of thy arg), SOME rhs] + |> Drule.instantiate' [SOME cty] [SOME (Thm.global_cterm_of thy arg), SOME rhs] |> Thm.varifyT_global end; @@ -207,10 +207,10 @@ val t = Thm.term_of ct; val T = fastype_of t; val mk_eq = - Thm.mk_binop (Proof_Context.cterm_of ctxt (Const (@{const_name Pure.eq}, T --> T --> propT))); + Thm.mk_binop (Thm.cterm_of ctxt (Const (@{const_name Pure.eq}, T --> T --> propT))); in case value ctxt t of NONE => Thm.reflexive ct - | SOME t' => conv ctxt (mk_eq ct (Proof_Context.cterm_of ctxt t')) RS @{thm eq_eq_TrueD} + | SOME t' => conv ctxt (mk_eq ct (Thm.cterm_of ctxt t')) RS @{thm eq_eq_TrueD} handle THM _ => error ("Failed to certify evaluation result of " ^ Syntax.string_of_term ctxt t) end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/coinduction.ML --- a/src/HOL/Tools/coinduction.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/coinduction.ML Fri Mar 06 15:58:56 2015 +0100 @@ -79,7 +79,7 @@ |> Ctr_Sugar_Util.list_exists_free vars |> Term.map_abs_vars (Variable.revert_fixed ctxt) |> fold_rev Term.absfree (names ~~ Ts) - |> Proof_Context.cterm_of ctxt; + |> Thm.cterm_of ctxt; val thm = Ctr_Sugar_Util.cterm_instantiate_pos [SOME phi] raw_thm; val e = length eqs; val p = length prems; @@ -88,7 +88,7 @@ EVERY' (map (fn var => resolve_tac ctxt [Ctr_Sugar_Util.cterm_instantiate_pos - [NONE, SOME (Proof_Context.cterm_of ctxt var)] exI]) vars), + [NONE, SOME (Thm.cterm_of ctxt var)] exI]) vars), if p = 0 then Ctr_Sugar_Util.CONJ_WRAP' (K (resolve_tac ctxt [refl])) eqs else REPEAT_DETERM_N e o (resolve_tac ctxt [conjI] THEN' resolve_tac ctxt [refl]) THEN' diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/datatype_realizer.ML --- a/src/HOL/Tools/datatype_realizer.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/datatype_realizer.ML Fri Mar 06 15:58:56 2015 +0100 @@ -105,11 +105,11 @@ (map (fn ((((i, _), T), U), tname) => make_pred i U T (mk_proj i is r) (Free (tname, T))) (descr ~~ recTs ~~ rec_result_Ts ~~ tnames))); - val inst = map (apply2 (Thm.cterm_of thy)) (map head_of (HOLogic.dest_conj + val inst = map (apply2 (Thm.global_cterm_of thy)) (map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct))) ~~ ps); val thm = - Goal.prove_internal ctxt (map (Thm.cterm_of thy) prems) (Thm.cterm_of thy concl) + Goal.prove_internal ctxt (map (Thm.global_cterm_of thy) prems) (Thm.global_cterm_of thy concl) (fn prems => EVERY [ rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]), @@ -185,12 +185,12 @@ val y' = Free ("y", T); val thm = - Goal.prove_internal ctxt (map (Thm.cterm_of thy) prems) - (Thm.cterm_of thy + Goal.prove_internal ctxt (map (Thm.global_cterm_of thy) prems) + (Thm.global_cterm_of thy (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y'])))) (fn prems => EVERY [ - rtac (cterm_instantiate [apply2 (Thm.cterm_of thy) (y, y')] exhaust) 1, + rtac (cterm_instantiate [apply2 (Thm.global_cterm_of thy) (y, y')] exhaust) 1, ALLGOALS (EVERY' [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites), resolve_tac ctxt prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])]) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/inductive.ML Fri Mar 06 15:58:56 2015 +0100 @@ -573,7 +573,7 @@ val elims = Induct.find_casesP ctxt prop; - val cprop = Thm.cterm_of thy prop; + val cprop = Thm.global_cterm_of thy prop; fun mk_elim rl = Thm.implies_intr cprop (Tactic.rule_by_tactic ctxt (mk_cases_tac ctxt) (Thm.assume cprop RS rl)) @@ -636,7 +636,7 @@ | _ => error ("equations matching pattern " ^ Syntax.string_of_term ctxt prop ^ " is not unique")); val inst = - map (fn v => (Thm.cterm_of thy (Var v), Thm.cterm_of thy (Envir.subst_term subst (Var v)))) + map (fn v => (Thm.global_cterm_of thy (Var v), Thm.global_cterm_of thy (Envir.subst_term subst (Var v)))) (Term.add_vars (lhs_of eq) []); in Drule.cterm_instantiate inst eq @@ -849,7 +849,7 @@ ||> is_auxiliary ? Local_Theory.restore_naming lthy; val fp_def' = Simplifier.rewrite (put_simpset HOL_basic_ss lthy' addsimps [fp_def]) - (Proof_Context.cterm_of lthy' (list_comb (rec_const, params))); + (Thm.cterm_of lthy' (list_comb (rec_const, params))); val specs = if length cs < 2 then [] else diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/inductive_set.ML Fri Mar 06 15:58:56 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 Thm.cterm_of thy o Var) vs)) + in Drule.instantiate' [] (rev (map (SOME o Thm.global_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 Thm.cterm_of thy) - rews ts) (Thm.reflexive (Thm.cterm_of thy h))) + (map2 (fn SOME r => K r | NONE => Thm.reflexive o Thm.global_cterm_of thy) + rews ts) (Thm.reflexive (Thm.global_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 - (Thm.cterm_of (Thm.theory_of_cterm ct) (eta_contract p (Thm.term_of ct))))))); + (Thm.global_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 - (Thm.cterm_of thy x, - Thm.cterm_of thy (fold_rev (Term.abs o pair "x") Ts + (Thm.global_cterm_of thy x, + Thm.global_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)))))) @@ -233,7 +233,7 @@ Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb in thm' RS (Drule.cterm_instantiate [(arg_cong_f, - Thm.cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT, + Thm.global_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) @@ -370,8 +370,8 @@ val T = HOLogic.mk_ptupleT ps Us; val x' = map_type (K (Rs ---> HOLogic.mk_setT T)) x in - (Thm.cterm_of thy x, - Thm.cterm_of thy (fold_rev (Term.abs o pair "x") Ts + (Thm.global_cterm_of thy x, + Thm.global_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; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/int_arith.ML --- a/src/HOL/Tools/int_arith.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/int_arith.ML Fri Mar 06 15:58:56 2015 +0100 @@ -82,7 +82,7 @@ fun number_of thy T n = if not (Sign.of_sort thy (T, @{sort numeral})) then raise CTERM ("number_of", []) - else Numeral.mk_cnumber (Thm.ctyp_of thy T) n; + else Numeral.mk_cnumber (Thm.global_ctyp_of thy T) n; val setup = Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2] diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/legacy_transfer.ML --- a/src/HOL/Tools/legacy_transfer.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/legacy_transfer.ML Fri Mar 06 15:58:56 2015 +0100 @@ -59,8 +59,8 @@ fun get_by_direction context (a, D) = let val ctxt = Context.proof_of context; - val a0 = Proof_Context.cterm_of ctxt a; - val D0 = Proof_Context.cterm_of ctxt D; + val a0 = Thm.cterm_of ctxt a; + val D0 = Thm.cterm_of ctxt D; fun eq_direction ((a, D), thm') = let val (a', D') = direction_of thm'; @@ -118,9 +118,9 @@ |> 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; + val c_vars = map (Thm.cterm_of ctxt3 o Var) vars; val (vs', ctxt4) = Variable.variant_fixes (map (fst o fst) vars) ctxt3; - val c_vars' = map (Proof_Context.cterm_of ctxt3 o (fn v => Free (v, bT))) vs'; + val c_vars' = map (Thm.cterm_of ctxt3 o (fn v => Free (v, bT))) vs'; val c_exprs' = map (Thm.apply a) c_vars'; (* transfer *) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/lin_arith.ML Fri Mar 06 15:58:56 2015 +0100 @@ -62,8 +62,8 @@ fun mk_nat_thm thy t = let - val cn = Thm.cterm_of thy (Var (("n", 0), HOLogic.natT)) - and ct = Thm.cterm_of thy t + val cn = Thm.global_cterm_of thy (Var (("n", 0), HOLogic.natT)) + and ct = Thm.global_cterm_of thy t in Drule.instantiate_normalize ([], [(cn, ct)]) @{thm le0} end; end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/monomorph.ML --- a/src/HOL/Tools/monomorph.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/monomorph.ML Fri Mar 06 15:58:56 2015 +0100 @@ -159,7 +159,7 @@ fun instantiate thy subst = let - fun cert (ix, (S, T)) = apply2 (Thm.ctyp_of thy) (TVar (ix, S), T) + fun cert (ix, (S, T)) = apply2 (Thm.global_ctyp_of thy) (TVar (ix, S), T) fun cert' subst = Vartab.fold (cons o cert) subst [] in Thm.instantiate (cert' subst, []) end diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/numeral.ML --- a/src/HOL/Tools/numeral.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/numeral.ML Fri Mar 06 15:58:56 2015 +0100 @@ -44,10 +44,10 @@ val oneT = Thm.ctyp_of_cterm one; val numeral = @{cpat "numeral"}; -val numeralT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of_cterm numeral)); +val numeralT = Thm.global_ctyp_of @{theory} (Term.range_type (Thm.typ_of_cterm numeral)); val uminus = @{cpat "uminus"}; -val uminusT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of_cterm uminus)); +val uminusT = Thm.global_ctyp_of @{theory} (Term.range_type (Thm.typ_of_cterm uminus)); fun instT T V = Thm.instantiate_cterm ([(V, T)], []); diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/numeral_simprocs.ML Fri Mar 06 15:58:56 2015 +0100 @@ -483,7 +483,7 @@ simplify_one ctxt (([th, cancel_th]) MRS trans); local - val Tp_Eq = Thm.reflexive (Thm.cterm_of @{theory HOL} HOLogic.Trueprop) + val Tp_Eq = Thm.reflexive (Thm.global_cterm_of @{theory HOL} HOLogic.Trueprop) fun Eq_True_elim Eq = Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI} in @@ -494,7 +494,7 @@ val pos = less $ zero $ t and neg = less $ t $ zero val thy = Proof_Context.theory_of ctxt fun prove p = - SOME (Eq_True_elim (Simplifier.asm_rewrite ctxt (Thm.cterm_of thy p))) + SOME (Eq_True_elim (Simplifier.asm_rewrite ctxt (Thm.global_cterm_of thy p))) handle THM _ => NONE in case prove pos of SOME th => SOME(th RS pos_th) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/record.ML Fri Mar 06 15:58:56 2015 +0100 @@ -78,7 +78,7 @@ fun match name ((name', _), _) = name = name'; fun getvar name = (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) + SOME var => Thm.global_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 (Thm.ctyp_of thy (Logic.varifyT_global rep_type))] []; + |> Drule.instantiate' [SOME (Thm.global_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,7 +141,7 @@ (*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", Thm.cterm_of typ_thy absC)] iso_tuple_intro; + rep_inject RS named_cterm_instantiate [("abst", Thm.global_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; @@ -1121,7 +1121,7 @@ fun get_upd_acc_cong_thm upd acc thy ss = let - val insts = [("upd", Thm.cterm_of thy upd), ("ac", Thm.cterm_of thy acc)]; + val insts = [("upd", Thm.global_cterm_of thy upd), ("ac", Thm.global_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 @@ -1375,9 +1375,9 @@ fun mk_split_free_tac free induct_thm i = let - val cfree = Thm.cterm_of thy free; + val cfree = Thm.global_cterm_of thy free; val _$ (_ $ r) = Thm.concl_of induct_thm; - val crec = Thm.cterm_of thy r; + val crec = Thm.global_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 @@ -1473,7 +1473,7 @@ | [x] => (head_of x, false)); val rule'' = cterm_instantiate - (map (apply2 (Proof_Context.cterm_of ctxt)) + (map (apply2 (Thm.cterm_of ctxt)) (case rev params of [] => (case AList.lookup (op =) (Term.add_frees g []) s of @@ -1606,7 +1606,7 @@ (roughly) the definition of the accessor.*) val surject = timeit_msg ctxt "record extension surjective proof:" (fn () => let - val cterm_ext = Thm.cterm_of defs_thy ext; + val cterm_ext = Thm.global_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 @@ -1756,7 +1756,7 @@ fun mk_eq_refl thy = @{thm equal_refl} |> Thm.instantiate - ([apply2 (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], []) + ([apply2 (Thm.global_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 = @@ -1943,8 +1943,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 = Thm.cterm_of ext_thy r_rec0; - val cterm_vrs = Thm.cterm_of ext_thy r_rec0_Vars; + val cterm_rec = Thm.global_cterm_of ext_thy r_rec0; + val cterm_vrs = Thm.global_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; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/reification.ML --- a/src/HOL/Tools/reification.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/reification.ML Fri Mar 06 15:58:56 2015 +0100 @@ -27,7 +27,7 @@ let val ct = case some_t of NONE => Thm.dest_arg concl - | SOME t => Proof_Context.cterm_of ctxt t + | SOME t => Thm.cterm_of ctxt t val thm = conv ct; in if Thm.is_reflexive thm then no_tac @@ -83,7 +83,7 @@ (fn {context, prems, ...} => Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym; val (cong' :: vars') = - Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Proof_Context.cterm_of ctxt'') vs); + Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Thm.cterm_of ctxt'') vs); val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'; in (vs', cong') end; @@ -140,8 +140,8 @@ val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt; val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta); (* FIXME !? *) val x = Free (xn, xT); - val cx = Proof_Context.cterm_of ctxt' x; - val cta = Proof_Context.cterm_of ctxt' ta; + val cx = Thm.cterm_of ctxt' x; + val cta = Thm.cterm_of ctxt' ta; val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of NONE => error "tryabsdecomp: Type not found in the Environement" | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT, @@ -169,12 +169,12 @@ val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv); val (fts, its) = (map (snd o snd) fnvs, - map (fn ((vn, vi), (tT, t)) => apply2 (Thm.cterm_of thy) (Var ((vn, vi), tT), t)) invs); + map (fn ((vn, vi), (tT, t)) => apply2 (Thm.global_cterm_of thy) (Var ((vn, vi), tT), t)) invs); val ctyenv = - map (fn ((vn, vi), (s, ty)) => apply2 (Thm.ctyp_of thy) (TVar ((vn, vi), s), ty)) + map (fn ((vn, vi), (s, ty)) => apply2 (Thm.global_ctyp_of thy) (TVar ((vn, vi), s), ty)) (Vartab.dest tyenv); in - ((map (Thm.cterm_of thy) fts ~~ replicate (length fts) ctxt, + ((map (Thm.global_cterm_of thy) fts ~~ replicate (length fts) ctxt, apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds) end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds)) end; @@ -203,14 +203,14 @@ val sbst = Envir.subst_term (tyenv, tmenv); val sbsT = Envir.subst_type tyenv; val subst_ty = - map (fn (n, (s, t)) => apply2 (Thm.ctyp_of thy) (TVar (n, s), t)) (Vartab.dest tyenv) + map (fn (n, (s, t)) => apply2 (Thm.global_ctyp_of thy) (TVar (n, s), t)) (Vartab.dest tyenv) val tml = Vartab.dest tmenv; val (subst_ns, bds) = fold_map (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds => let val name = snd (the (AList.lookup (op =) tml xn0)); val (idx, bds) = index_of name bds; - in (apply2 (Thm.cterm_of thy) (n, idx |> HOLogic.mk_nat), bds) end) subst bds; + in (apply2 (Thm.global_cterm_of thy) (n, idx |> HOLogic.mk_nat), bds) end) subst bds; val subst_vs = let fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) = @@ -220,10 +220,10 @@ val (bsT, _) = the (AList.lookup Type.could_unify bds lT); val vsn = the (AList.lookup (op =) vsns_map vs); val vs' = fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')); - in apply2 (Thm.cterm_of thy) (vs, vs') end; + in apply2 (Thm.global_cterm_of thy) (vs, vs') end; in map h subst end; val cts = - map (fn ((vn, vi), (tT, t)) => apply2 (Thm.cterm_of thy) (Var ((vn, vi), tT), t)) + map (fn ((vn, vi), (tT, t)) => apply2 (Thm.global_cterm_of thy) (Var ((vn, vi), tT), t)) (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b)) (map (fn n => (n, 0)) xns) tml); val substt = @@ -261,14 +261,14 @@ val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt; val subst = the o AList.lookup (op =) - (map2 (fn T => fn v => (T, Proof_Context.cterm_of ctxt' (Free (v, T)))) tys vs); + (map2 (fn T => fn v => (T, Thm.cterm_of ctxt' (Free (v, T)))) tys vs); fun prep_eq eq = let 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 (Proof_Context.cterm_of ctxt' v, subst T) + (fn (v as Var (_, T)) => SOME (Thm.cterm_of ctxt' v, subst T) | _ => NONE) vs; in Thm.instantiate ([], subst) eq end; val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs; @@ -287,7 +287,7 @@ |> strip_comb |> snd |> filter is_list_var; val vs = map (fn v as Var (_, T) => (v, the (AList.lookup Type.could_unify bds' T) |> snd |> HOLogic.mk_list (dest_listT T))) vars; - val th' = Drule.instantiate_normalize ([], map (apply2 (Proof_Context.cterm_of ctxt)) vs) th; + val th' = Drule.instantiate_normalize ([], map (apply2 (Thm.cterm_of ctxt)) vs) th; val th'' = Thm.symmetric (dereify ctxt [] (Thm.lhs_of th')); in Thm.transitive th'' th' end; diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/sat.ML --- a/src/HOL/Tools/sat.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/sat.ML Fri Mar 06 15:58:56 2015 +0100 @@ -73,7 +73,7 @@ val resolution_thm = @{lemma "(P ==> False) ==> (~ P ==> False) ==> False" by (rule case_split)} -val cP = Thm.cterm_of @{theory} (Var (("P", 0), HOLogic.boolT)); +val cP = Thm.global_cterm_of @{theory} (Var (("P", 0), HOLogic.boolT)); (* ------------------------------------------------------------------------- *) (* lit_ord: an order on integers that considers their absolute values only, *) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/set_comprehension_pointfree.ML --- a/src/HOL/Tools/set_comprehension_pointfree.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/set_comprehension_pointfree.ML Fri Mar 06 15:58:56 2015 +0100 @@ -468,7 +468,7 @@ fun conv ctxt t = let val ([t'], ctxt') = Variable.import_terms true [t] (Variable.declare_term t ctxt) - val ct = Proof_Context.cterm_of ctxt' t' + val ct = Thm.cterm_of ctxt' t' fun unfold_conv thms = Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) (empty_simpset ctxt' addsimps thms) @@ -498,7 +498,7 @@ val arg_cong = Thm.incr_indexes (maxidx_of_term pred + 1) @{thm 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 + cterm_instantiate [(Thm.cterm_of ctxt f, Thm.cterm_of ctxt pred)] arg_cong end; fun simproc ctxt redex = diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Tools/split_rule.ML Fri Mar 06 15:58:56 2015 +0100 @@ -42,7 +42,7 @@ fun split_rule_var' (t as Var (v, Type ("fun", [T1, T2]))) rl = let val T' = HOLogic.flatten_tupleT T1 ---> T2; val newt = ap_split T1 T2 (Var (v, T')); - val cterm = Thm.cterm_of (Thm.theory_of_thm rl); + val cterm = Thm.global_cterm_of (Thm.theory_of_thm rl); in Thm.instantiate ([], [(cterm t, cterm newt)]) rl end | split_rule_var' _ rl = rl; @@ -56,7 +56,7 @@ fun complete_split_rule_var (t as Var (v, T), ts) (rl, vs) = let - val cterm = Thm.cterm_of (Thm.theory_of_thm rl) + val cterm = Thm.global_cterm_of (Thm.theory_of_thm rl) val (Us', U') = strip_type T; val Us = take (length ts) Us'; val U = drop (length ts) Us' ---> U'; @@ -70,7 +70,7 @@ end | mk_tuple _ x = x; val newt = ap_split' Us U (Var (v, T')); - val cterm = Thm.cterm_of (Thm.theory_of_thm rl); + val cterm = Thm.global_cterm_of (Thm.theory_of_thm rl); val (vs', insts) = fold mk_tuple ts (vs, []); in (Drule.instantiate_normalize ([], [(cterm t, cterm newt)] @ insts) rl, vs') diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/Word/WordBitwise.thy Fri Mar 06 15:58:56 2015 +0100 @@ -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 |> Thm.cterm_of thy; + |> HOLogic.mk_Trueprop |> Thm.global_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)) diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/ex/Iff_Oracle.thy --- a/src/HOL/ex/Iff_Oracle.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/ex/Iff_Oracle.thy Fri Mar 06 15:58:56 2015 +0100 @@ -24,7 +24,7 @@ in fn (thy, n) => if n > 0 andalso n mod 2 = 0 - then Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_iff n)) + then Thm.global_cterm_of thy (HOLogic.mk_Trueprop (mk_iff n)) else raise Fail ("iff_oracle: " ^ string_of_int n) end \ diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/ex/SAT_Examples.thy --- a/src/HOL/ex/SAT_Examples.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/ex/SAT_Examples.thy Fri Mar 06 15:58:56 2015 +0100 @@ -534,7 +534,7 @@ | and_to_list fm acc = rev (fm :: acc) val clauses = and_to_list prop_fm [] val terms = map (HOLogic.mk_Trueprop o Prop_Logic.term_of_prop_formula) clauses - val cterms = map (Thm.cterm_of @{theory}) terms + val cterms = map (Thm.global_cterm_of @{theory}) terms val start = Timing.start () val _ = SAT.rawsat_thm @{context} cterms in diff -r 92d7d8e4f1bf -r 291934bac95e src/HOL/ex/SVC_Oracle.thy --- a/src/HOL/ex/SVC_Oracle.thy Fri Mar 06 14:01:08 2015 +0100 +++ b/src/HOL/ex/SVC_Oracle.thy Fri Mar 06 15:58:56 2015 +0100 @@ -113,7 +113,7 @@ let val thy = Thm.theory_of_cterm ct; val (abs_goal, _) = svc_abstract (Thm.term_of ct); - val th = svc_oracle (Thm.cterm_of thy abs_goal); + val th = svc_oracle (Thm.global_cterm_of thy abs_goal); in compose_tac ctxt (false, th, 0) i end handle TERM _ => no_tac); *} diff -r 92d7d8e4f1bf -r 291934bac95e src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Provers/Arith/fast_lin_arith.ML Fri Mar 06 15:58:56 2015 +0100 @@ -905,7 +905,7 @@ let val thy = Proof_Context.theory_of ctxt val nTconcl = LA_Logic.neg_prop Tconcl - val cnTconcl = Thm.cterm_of thy nTconcl + val cnTconcl = Thm.global_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 diff -r 92d7d8e4f1bf -r 291934bac95e src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Provers/hypsubst.ML Fri Mar 06 15:58:56 2015 +0100 @@ -184,10 +184,10 @@ (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 (Thm.cterm_of thy o Logic.mk_type) (V, U)); + val (instT, _) = Thm.match (apply2 (Thm.global_cterm_of thy o Logic.mk_type) (V, U)); in compose_tac ctxt (true, Drule.instantiate_normalize (instT, - map (apply2 (Thm.cterm_of thy)) + map (apply2 (Thm.global_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', diff -r 92d7d8e4f1bf -r 291934bac95e src/Provers/splitter.ML --- a/src/Provers/splitter.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Provers/splitter.ML Fri Mar 06 15:58:56 2015 +0100 @@ -315,7 +315,7 @@ fun inst_lift Ts t (T, U, pos) state i = let - val cert = Thm.cterm_of (Thm.theory_of_thm state); + val cert = Thm.global_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); @@ -343,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 = Thm.cterm_of (Thm.theory_of_thm state); + val cert = Thm.global_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; diff -r 92d7d8e4f1bf -r 291934bac95e src/Provers/trancl.ML --- a/src/Provers/trancl.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Provers/trancl.ML Fri Mar 06 15:58:56 2015 +0100 @@ -94,7 +94,7 @@ let fun inst thm = 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; + in Drule.cterm_instantiate [(Thm.global_cterm_of thy r', Thm.global_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; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/class_declaration.ML --- a/src/Pure/Isar/class_declaration.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/class_declaration.ML Fri Mar 06 15:58:56 2015 +0100 @@ -93,7 +93,7 @@ val sup_of_classes = map (snd o Class.rules thy) sups; val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class); val axclass_intro = #intro (Axclass.get_info thy class); - val base_sort_trivs = Thm.of_sort (Thm.ctyp_of thy aT, base_sort); + val base_sort_trivs = Thm.of_sort (Thm.global_ctyp_of thy aT, base_sort); fun tac ctxt = REPEAT (SOMEGOAL (match_tac ctxt (axclass_intro :: sup_of_classes @ loc_axiom_intros @ base_sort_trivs) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/code.ML Fri Mar 06 15:58:56 2015 +0100 @@ -363,7 +363,7 @@ fun analyze_constructor thy (c, ty) = let - val _ = Thm.cterm_of thy (Const (c, ty)); + val _ = Thm.global_cterm_of thy (Const (c, ty)); val ty_decl = devarify (const_typ thy c); fun last_typ c_ty ty = let @@ -611,7 +611,7 @@ val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) raw_vars; fun expand (v, ty) thm = Drule.fun_cong_rule thm - (Thm.cterm_of thy (Var ((v, 0), ty))); + (Thm.global_cterm_of thy (Var ((v, 0), ty))); in thm |> fold expand vars @@ -681,7 +681,7 @@ (* code equation certificates *) fun build_head thy (c, ty) = - Thm.cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); + Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); fun get_head thy cert_thm = let @@ -716,7 +716,7 @@ val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm val ty = fastype_of lhs; val ty_abs = (fastype_of o snd o dest_comb) lhs; - val abs = Thm.cterm_of thy (Const (c, ty --> ty_abs)); + val abs = Thm.global_cterm_of thy (Const (c, ty --> ty_abs)); val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric cert, Thm.combination (Thm.reflexive abs) abs_thm]; in (c, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; @@ -804,7 +804,7 @@ |> constrain_thm thy vs sorts; val head' = Thm.term_of head |> subst - |> Thm.cterm_of thy; + |> Thm.global_cterm_of thy; val cert_thm'' = cert_thm' |> Thm.elim_implies (Thm.assume head'); in cert_thm'' end; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/element.ML Fri Mar 06 15:58:56 2015 +0100 @@ -203,7 +203,7 @@ let val thy = Thm.theory_of_thm th; val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C); - val th' = Thm.instantiate ([], [(Thm.cterm_of thy C, Thm.cterm_of thy thesis)]) th; + val th' = Thm.instantiate ([], [(Thm.global_cterm_of thy C, Thm.global_cterm_of thy thesis)]) th; in (th', true) end | NONE => (th, false)); @@ -339,7 +339,7 @@ fun instantiate_tfrees thy subst th = let val idx = Thm.maxidx_of th + 1; - fun cert_inst (a, (S, T)) = (Thm.ctyp_of thy (TVar ((a, idx), S)), Thm.ctyp_of thy T); + fun cert_inst (a, (S, T)) = (Thm.global_ctyp_of thy (TVar ((a, idx), S)), Thm.global_ctyp_of thy T); fun add_inst (a, S) insts = if AList.defined (op =) insts a then insts @@ -354,8 +354,8 @@ end; fun instantiate_frees thy subst = - Drule.forall_intr_list (map (Thm.cterm_of thy o Free o fst) subst) #> - Drule.forall_elim_list (map (Thm.cterm_of thy o snd) subst); + Drule.forall_intr_list (map (Thm.global_cterm_of thy o Free o fst) subst) #> + Drule.forall_elim_list (map (Thm.global_cterm_of thy o snd) subst); fun hyps_rule rule th = let val {hyps, ...} = Thm.crep_thm th in diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/expression.ML Fri Mar 06 15:58:56 2015 +0100 @@ -629,8 +629,8 @@ val bodyT = Term.fastype_of body; in if bodyT = propT - then (t, propT, Thm.reflexive (Thm.cterm_of thy t)) - else (body, bodyT, Object_Logic.atomize (Proof_Context.init_global thy) (Thm.cterm_of thy t)) + then (t, propT, Thm.reflexive (Thm.global_cterm_of thy t)) + else (body, bodyT, Object_Logic.atomize (Proof_Context.init_global thy) (Thm.global_cterm_of thy t)) end; (* achieve plain syntax for locale predicates (without "PROP") *) @@ -684,11 +684,11 @@ compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN compose_tac defs_ctxt (false, - Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_thy) norm_ts), 0) 1); + Conjunction.intr_balanced (map (Thm.assume o Thm.global_cterm_of defs_thy) norm_ts), 0) 1); val conjuncts = (Drule.equal_elim_rule2 OF - [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_thy statement))]) + [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.global_cterm_of defs_thy statement))]) |> Conjunction.elim_balanced (length ts); val (_, axioms_ctxt) = defs_ctxt @@ -705,7 +705,7 @@ fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val ctxt = Proof_Context.init_global thy; - val defs' = map (Thm.cterm_of thy #> Assumption.assume ctxt #> Drule.abs_def) defs; + val defs' = map (Thm.global_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) @@ -757,7 +757,7 @@ fun defines_to_notes ctxt (Defines defs) = Notes ("", map (fn (a, (def, _)) => - (a, [([Assumption.assume ctxt (Proof_Context.cterm_of ctxt def)], + (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], [(Attrib.internal o K) Locale.witness_add])])) defs) | defines_to_notes _ e = e; @@ -806,7 +806,7 @@ val notes = if is_some asm then [("", [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []), - [([Assumption.assume pred_ctxt (Thm.cterm_of thy' (the asm))], + [([Assumption.assume pred_ctxt (Thm.global_cterm_of thy' (the asm))], [(Attrib.internal o K) Locale.witness_add])])])] else []; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/generic_target.ML Fri Mar 06 15:58:56 2015 +0100 @@ -179,7 +179,7 @@ val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) - val ((defs, _), rhs') = Proof_Context.cterm_of lthy rhs + val ((defs, _), rhs') = Thm.cterm_of lthy rhs |> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; val xs = Variable.add_fixed lthy rhs' []; @@ -205,7 +205,7 @@ (*result*) val def = Thm.transitive local_def global_def - |> Local_Defs.contract lthy3 defs (Proof_Context.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); + |> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); val ([(res_name, [res])], lthy4) = lthy3 |> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; in ((lhs, (res_name, res)), lthy4) end; @@ -228,7 +228,7 @@ val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []); val frees = map Free (Thm.fold_terms Term.add_frees th' []); val (th'' :: vs) = - (th' :: map (Drule.mk_term o Proof_Context.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) + (th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; @@ -243,9 +243,9 @@ val instT = map_filter (fn (TVar v, T) => SOME (v, T) | _ => NONE) (tvars ~~ tfrees); val inst = filter (is_Var o fst) (vars ~~ frees); val cinstT = instT - |> map (apply2 (Proof_Context.ctyp_of ctxt) o apfst TVar); + |> map (apply2 (Thm.ctyp_of ctxt) o apfst TVar); val cinst = inst - |> map (apply2 (Proof_Context.cterm_of ctxt o Term.map_types (Term_Subst.instantiateT instT))); + |> map (apply2 (Thm.cterm_of ctxt o Term.map_types (Term_Subst.instantiateT instT))); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/local_defs.ML --- a/src/Pure/Isar/local_defs.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/local_defs.ML Fri Mar 06 15:58:56 2015 +0100 @@ -214,7 +214,7 @@ fun derived_def ctxt conditional prop = let val ((c, T), rhs) = prop - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt |> meta_rewrite_conv ctxt |> (snd o Logic.dest_equals o Thm.prop_of) |> conditional ? Logic.strip_imp_concl diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/obtain.ML Fri Mar 06 15:58:56 2015 +0100 @@ -129,7 +129,7 @@ (*obtain parms*) val (Ts, parms_ctxt) = fold_map Proof_Context.inferred_param xs' asms_ctxt; val parms = map Free (xs' ~~ Ts); - val cparms = map (Proof_Context.cterm_of ctxt) parms; + val cparms = map (Thm.cterm_of ctxt) parms; val _ = Variable.warn_extra_tfrees fix_ctxt parms_ctxt; (*obtain statements*) @@ -187,17 +187,17 @@ fun result tac facts ctxt = let val ((thesis_var, thesis), thesis_ctxt) = bind_judgment ctxt Auto_Bind.thesisN; - val st = Goal.init (Proof_Context.cterm_of ctxt thesis); + val st = Goal.init (Thm.cterm_of ctxt thesis); val rule = (case SINGLE (Method.insert_tac facts 1 THEN tac thesis_ctxt) st of NONE => raise THM ("Obtain.result: tactic failed", 0, facts) | SOME th => check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th))); - val closed_rule = Thm.forall_intr (Proof_Context.cterm_of ctxt (Free thesis_var)) rule; + val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule; val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; val obtain_rule = - Thm.forall_elim (Proof_Context.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; + Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; val ((params, stmt), fix_ctxt) = Variable.focus_cterm (Thm.cprem_of obtain_rule 1) ctxt'; val (prems, ctxt'') = Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) @@ -237,20 +237,20 @@ val xs = map (apsnd norm_type o fst) vars; val ys = map (apsnd norm_type) (drop m params); val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; - val terms = map (Drule.mk_term o Thm.cterm_of thy o Free) (xs @ ys'); + val terms = map (Drule.mk_term o Thm.global_cterm_of thy o Free) (xs @ ys'); val instT = fold (Term.add_tvarsT o #2) params [] - |> map (TVar #> (fn T => (Thm.ctyp_of thy T, Thm.ctyp_of thy (norm_type T)))); + |> map (TVar #> (fn T => (Thm.global_ctyp_of thy T, Thm.global_ctyp_of thy (norm_type T)))); val closed_rule = rule - |> Thm.forall_intr (Thm.cterm_of thy (Free thesis_var)) + |> Thm.forall_intr (Thm.global_cterm_of thy (Free thesis_var)) |> Thm.instantiate (instT, []); val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt; val vars' = map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~ (map snd vars @ replicate (length ys) NoSyn); - val rule'' = Thm.forall_elim (Thm.cterm_of thy (Logic.varify_global (Free thesis_var))) rule'; + val rule'' = Thm.forall_elim (Thm.global_cterm_of thy (Logic.varify_global (Free thesis_var))) rule'; in ((vars', rule''), ctxt') end; fun inferred_type (binding, _, mx) ctxt = @@ -288,7 +288,7 @@ |> Proof.map_context (K ctxt') |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms) |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm - (obtain_export fix_ctxt rule (map (Proof_Context.cterm_of ctxt) ts)) + (obtain_export fix_ctxt rule (map (Thm.cterm_of ctxt) ts)) [(Thm.empty_binding, asms)]) |> Proof.bind_terms Auto_Bind.no_facts end; @@ -312,7 +312,7 @@ |> Proof.local_goal print_result (K I) (pair o rpair I) "guess" (SOME before_qed) after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])] |> Proof.refine (Method.primitive_text (fn _ => fn _ => - Goal.init (Proof_Context.cterm_of ctxt thesis))) |> Seq.hd + Goal.init (Thm.cterm_of ctxt thesis))) |> Seq.hd end; in diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/proof.ML Fri Mar 06 15:58:56 2015 +0100 @@ -910,7 +910,7 @@ val goal_propss = filter_out null propss'; val goal = Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss) - |> Proof_Context.cterm_of ctxt + |> Thm.cterm_of ctxt |> Thm.weaken_sorts (Variable.sorts_of (context_of goal_state)); val statement = ((kind, pos), propss', Thm.term_of goal); val after_qed' = after_qed |>> (fn after_local => diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/proof_context.ML Fri Mar 06 15:58:56 2015 +0100 @@ -11,8 +11,6 @@ val theory_of: Proof.context -> theory val init_global: theory -> Proof.context val get_global: theory -> string -> Proof.context - val cterm_of: Proof.context -> term -> cterm - val ctyp_of: Proof.context -> typ -> ctyp type mode val mode_default: mode val mode_stmt: mode @@ -182,9 +180,6 @@ val init_global = Proof_Context.init_global; val get_global = Proof_Context.get_global; -val cterm_of = Thm.cterm_of o theory_of; -val ctyp_of = Thm.ctyp_of o theory_of; - (** inner syntax mode **) @@ -1169,7 +1164,7 @@ val ((propss, _), ctxt1) = prepp (map snd args) ctxt; val _ = Variable.warn_extra_tfrees ctxt ctxt1; val (premss, ctxt2) = - fold_burrow (Assumption.add_assms exp o map (cterm_of ctxt)) propss ctxt1; + fold_burrow (Assumption.add_assms exp o map (Thm.cterm_of ctxt)) propss ctxt1; in ctxt2 |> auto_bind_facts (flat propss) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Isar/spec_rules.ML --- a/src/Pure/Isar/spec_rules.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Isar/spec_rules.ML Fri Mar 06 15:58:56 2015 +0100 @@ -48,7 +48,7 @@ fun add class (ts, ths) lthy = let - val cts = map (Proof_Context.cterm_of lthy) ts; + val cts = map (Thm.cterm_of lthy) ts; in lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/ML/ml_antiquotations.ML --- a/src/Pure/ML/ml_antiquotations.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/ML/ml_antiquotations.ML Fri Mar 06 15:58:56 2015 +0100 @@ -62,18 +62,18 @@ ML_Antiquotation.inline @{binding prop} (Args.prop >> (ML_Syntax.atomic o ML_Syntax.print_term)) #> ML_Antiquotation.value @{binding ctyp} (Args.typ >> (fn T => - "Proof_Context.ctyp_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_typ T))) #> + "Thm.ctyp_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_typ T))) #> ML_Antiquotation.value @{binding cterm} (Args.term >> (fn t => - "Proof_Context.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> + "Thm.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> ML_Antiquotation.value @{binding cprop} (Args.prop >> (fn t => - "Proof_Context.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> + "Thm.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t))) #> ML_Antiquotation.value @{binding cpat} (Args.context -- Scan.lift Args.name_inner_syntax >> uncurry Proof_Context.read_term_pattern >> (fn t => - "Proof_Context.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t)))); + "Thm.cterm_of ML_context " ^ ML_Syntax.atomic (ML_Syntax.print_term t)))); (* type classes *) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Proof/proof_checker.ML --- a/src/Pure/Proof/proof_checker.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Proof/proof_checker.ML Fri Mar 06 15:58:56 2015 +0100 @@ -78,7 +78,7 @@ let val tvars = Term.add_tvars (Thm.full_prop_of thm) [] |> rev; val (fmap, thm') = Thm.varifyT_global' [] thm; - val ctye = map (apply2 (Thm.ctyp_of thy)) + val ctye = map (apply2 (Thm.global_ctyp_of thy)) (map TVar tvars @ map (fn ((_, S), ixn) => TVar (ixn, S)) fmap ~~ Ts) in Thm.instantiate (ctye, []) (forall_intr_vars (Thm.forall_intr_frees thm')) @@ -106,13 +106,13 @@ val (x, names') = Name.variant s names; val thm = thm_of ((x, T) :: vs, names') Hs prf in - Thm.forall_intr (Thm.cterm_of thy (Free (x, T))) thm + Thm.forall_intr (Thm.global_cterm_of thy (Free (x, T))) thm end | thm_of (vs, names) Hs (prf % SOME t) = let val thm = thm_of (vs, names) Hs prf; - val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t)); + val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); in Thm.forall_elim ct thm handle THM (s, _, _) => appl_error thy pretty_term vs Hs s prf t @@ -120,7 +120,7 @@ | thm_of (vs, names) Hs (AbsP (s, SOME t, prf)) = let - val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t)); + val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf; in Thm.implies_intr ct thm @@ -135,7 +135,7 @@ handle THM (s, _, _) => appl_error thy pretty_prf (fst vars) Hs s prf prf' end - | thm_of _ _ (Hyp t) = Thm.assume (Thm.cterm_of thy t) + | thm_of _ _ (Hyp t) = Thm.assume (Thm.global_cterm_of thy t) | thm_of _ _ _ = error "thm_of_proof: partial proof term"; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Proof/proof_syntax.ML Fri Mar 06 15:58:56 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 - (Thm.cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of) + (Thm.global_cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of) end; fun read_term thy topsort = diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/Tools/rule_insts.ML --- a/src/Pure/Tools/rule_insts.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/Tools/rule_insts.ML Fri Mar 06 15:58:56 2015 +0100 @@ -116,8 +116,8 @@ val inst_tvars = map_filter (make_instT (instT2 o instT1)) tvars; val inst_vars = map_filter (make_inst inst2) vars2; in - (map (apply2 (Thm.ctyp_of thy)) inst_tvars, - map (apply2 (Thm.cterm_of thy)) inst_vars) + (map (apply2 (Thm.global_ctyp_of thy)) inst_tvars, + map (apply2 (Thm.global_cterm_of thy)) inst_vars) end; fun where_rule ctxt mixed_insts fixes thm = @@ -250,7 +250,7 @@ val cenv = map (fn (xi, t) => - apply2 (Thm.cterm_of thy) (Var (xi, fastype_of t), t)) + apply2 (Thm.global_cterm_of thy) (Var (xi, fastype_of t), t)) (distinct (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) (xis ~~ ts)); @@ -264,7 +264,7 @@ 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 (Thm.ctyp_of thy o Logic.incr_tvar inc); + val lifttvar = apply2 (Thm.global_ctyp_of thy o Logic.incr_tvar inc); val rule = Drule.instantiate_normalize (map lifttvar envT', map liftpair cenv) (Thm.lift_rule cgoal thm) @@ -283,7 +283,7 @@ let val thy = Thm.theory_of_thm rl; val maxidx = Thm.maxidx_of rl; - fun cvar xi = Thm.cterm_of thy (Var (xi, propT)); + fun cvar xi = Thm.global_cterm_of thy (Var (xi, propT)); val revcut_rl' = Drule.instantiate_normalize ([], [(cvar ("V", 0), cvar ("V", maxidx + 1)), (cvar ("W", 0), cvar ("W", maxidx + 1))]) Drule.revcut_rl; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/axclass.ML --- a/src/Pure/axclass.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/axclass.ML Fri Mar 06 15:58:56 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 (Thm.ctyp_of thy2 (TVar ((Name.aT, 0), [])))] [] + |> Drule.instantiate' [SOME (Thm.global_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 (Thm.ctyp_of thy (TVar ((a, 0), [])))) names; + val std_vars = map (fn a => SOME (Thm.global_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 (Thm.ctyp_of thy' (TVar ((Name.aT, 0), [])))] []; + |> Drule.instantiate' [SOME (Thm.global_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 (Thm.ctyp_of thy' (TVar ((a, 0), [])))) args; + val std_vars = map (fn (a, S) => SOME (Thm.global_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')) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/conjunction.ML --- a/src/Pure/conjunction.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/conjunction.ML Fri Mar 06 15:58:56 2015 +0100 @@ -29,7 +29,7 @@ (** abstract syntax **) -fun certify t = Thm.cterm_of (Context.the_theory (Context.the_thread_data ())) t; +fun certify t = Thm.global_cterm_of (Context.the_theory (Context.the_thread_data ())) t; val read_prop = certify o Simple_Syntax.read_prop; val true_prop = certify Logic.true_prop; @@ -133,7 +133,7 @@ local fun conjs thy n = - let val As = map (fn A => Thm.cterm_of thy (Free (A, propT))) (Name.invent Name.context "A" n) + let val As = map (fn A => Thm.global_cterm_of thy (Free (A, propT))) (Name.invent Name.context "A" n) in (As, mk_conjunction_balanced As) end; val B = read_prop "B"; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/drule.ML --- a/src/Pure/drule.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/drule.ML Fri Mar 06 15:58:56 2015 +0100 @@ -132,10 +132,10 @@ (*The premises of a theorem, as a cterm list*) 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)); +fun cterm_fun f ct = Thm.global_cterm_of (Thm.theory_of_cterm ct) (f (Thm.term_of ct)); +fun ctyp_fun f cT = Thm.global_ctyp_of (Thm.theory_of_ctyp cT) (f (Thm.typ_of cT)); -fun certify t = Thm.cterm_of (Context.the_theory (Context.the_thread_data ())) t; +fun certify t = Thm.global_cterm_of (Context.the_theory (Context.the_thread_data ())) t; val implies = certify Logic.implies; fun mk_implies (A, B) = Thm.apply (Thm.apply implies A) B; @@ -201,7 +201,7 @@ (*Generalization over Vars -- canonical order*) fun forall_intr_vars th = fold Thm.forall_intr - (map (Thm.cterm_of (Thm.theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th [])) th; + (map (Thm.global_cterm_of (Thm.theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th [])) th; fun outer_params t = let val vs = Term.strip_all_vars t @@ -211,7 +211,7 @@ fun gen_all th = let val {thy, prop, maxidx, ...} = Thm.rep_thm th; - fun elim (x, T) = Thm.forall_elim (Thm.cterm_of thy (Var ((x, maxidx + 1), T))); + fun elim (x, T) = Thm.forall_elim (Thm.global_cterm_of thy (Var ((x, maxidx + 1), T))); in fold elim (outer_params prop) th end; (*lift vars wrt. outermost goal parameters @@ -224,10 +224,10 @@ |> map (fn (x, T) => Var ((x, maxidx + 1), Logic.incr_tvar (maxidx + 1) T)); val Ts = map Term.fastype_of ps; val inst = Thm.fold_terms Term.add_vars th [] |> map (fn (xi, T) => - (Thm.cterm_of thy (Var (xi, T)), Thm.cterm_of thy (Term.list_comb (Var (xi, Ts ---> T), ps)))); + (Thm.global_cterm_of thy (Var (xi, T)), Thm.global_cterm_of thy (Term.list_comb (Var (xi, Ts ---> T), ps)))); in th |> Thm.instantiate ([], inst) - |> fold_rev (Thm.forall_intr o Thm.cterm_of thy) ps + |> fold_rev (Thm.forall_intr o Thm.global_cterm_of thy) ps end; (*direct generalization*) @@ -248,8 +248,8 @@ let val thy = Theory.merge_list (map Thm.theory_of_thm ths); val (instT, inst) = Term_Subst.zero_var_indexes_inst (map Thm.full_prop_of ths); - val cinstT = map (fn (v, T) => (Thm.ctyp_of thy (TVar v), Thm.ctyp_of thy T)) instT; - val cinst = map (fn (v, t) => (Thm.cterm_of thy (Var v), Thm.cterm_of thy t)) inst; + val cinstT = map (fn (v, T) => (Thm.global_ctyp_of thy (TVar v), Thm.global_ctyp_of thy T)) instT; + val cinst = map (fn (v, t) => (Thm.global_cterm_of thy (Var v), Thm.global_cterm_of thy t)) inst; in map (Thm.adjust_maxidx_thm ~1 o Thm.instantiate (cinstT, cinst)) ths end; val zero_var_indexes = singleton zero_var_indexes_list; @@ -644,9 +644,9 @@ let val thy = Thm.theory_of_cterm ct; val T = Thm.typ_of_cterm ct; - val a = Thm.ctyp_of thy (TVar (("'a", 0), [])); - val x = Thm.cterm_of thy (Var (("x", 0), T)); - in Thm.instantiate ([(a, Thm.ctyp_of thy T)], [(x, ct)]) termI end; + val a = Thm.global_ctyp_of thy (TVar (("'a", 0), [])); + val x = Thm.global_cterm_of thy (Var (("x", 0), T)); + in Thm.instantiate ([(a, Thm.global_ctyp_of thy T)], [(x, ct)]) termI end; fun dest_term th = let val cprop = strip_imp_concl (Thm.cprop_of th) in @@ -789,7 +789,7 @@ val (thy, tye, _) = fold_rev add_types ctpairs (Thm.theory_of_thm th, Vartab.empty, 0); val instT = Vartab.fold (fn (xi, (S, T)) => - cons (Thm.ctyp_of thy (TVar (xi, S)), Thm.ctyp_of thy (Envir.norm_type tye T))) tye []; + cons (Thm.global_ctyp_of thy (TVar (xi, S)), Thm.global_ctyp_of thy (Envir.norm_type tye T))) tye []; val inst = map (apply2 (Thm.instantiate_cterm (instT, []))) ctpairs; in instantiate_normalize (instT, inst) th end handle TERM (msg, _) => raise THM (msg, 0, [th]) @@ -807,11 +807,11 @@ map_filter (Option.map Thm.term_of) cts); fun inst_of (v, ct) = - (Thm.cterm_of (Thm.theory_of_cterm ct) (Var v), ct) + (Thm.global_cterm_of (Thm.theory_of_cterm ct) (Var v), ct) handle TYPE (msg, _, _) => err msg; fun tyinst_of (v, cT) = - (Thm.ctyp_of (Thm.theory_of_ctyp cT) (TVar v), cT) + (Thm.global_ctyp_of (Thm.theory_of_ctyp cT) (TVar v), cT) handle TYPE (msg, _, _) => err msg; fun zip_vars xs ys = @@ -843,7 +843,7 @@ fun ren (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, ren t) | ren (t $ u) = ren t $ ren u | ren t = t; - in Thm.equal_elim (Thm.reflexive (Thm.cterm_of thy (ren (Thm.prop_of thm)))) thm end; + in Thm.equal_elim (Thm.reflexive (Thm.global_cterm_of thy (ren (Thm.prop_of thm)))) thm end; (* renaming in left-to-right order *) @@ -864,7 +864,7 @@ | rename xs t = (xs, t); in (case rename xs prop of - ([], prop') => Thm.equal_elim (Thm.reflexive (Thm.cterm_of thy prop')) thm + ([], prop') => Thm.equal_elim (Thm.reflexive (Thm.global_cterm_of thy prop')) thm | _ => error "More names than abstractions in theorem") end; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/goal.ML --- a/src/Pure/goal.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/goal.ML Fri Mar 06 15:58:56 2015 +0100 @@ -133,15 +133,15 @@ val As = map Thm.term_of assms; val xs = map Free (fold Term.add_frees (prop :: As) []); - val fixes = map (Thm.cterm_of thy) xs; + val fixes = map (Thm.global_cterm_of thy) xs; val tfrees = fold Term.add_tfrees (prop :: As) []; val instT = - map (fn (a, S) => (Thm.ctyp_of thy (TVar ((a, 0), S)), Thm.ctyp_of thy (TFree (a, S)))) tfrees; + map (fn (a, S) => (Thm.global_ctyp_of thy (TVar ((a, 0), S)), Thm.global_ctyp_of thy (TFree (a, S)))) tfrees; val global_prop = Logic.varify_types_global (fold_rev Logic.all xs (Logic.list_implies (As, prop))) - |> Thm.cterm_of thy + |> Thm.global_cterm_of thy |> Thm.weaken_sorts (Variable.sorts_of ctxt); val global_result = result |> Future.map (Drule.flexflex_unique (SOME ctxt) #> @@ -191,7 +191,7 @@ ("The error(s) above occurred for the goal statement" ^ Position.here pos ^ ":\n" ^ Syntax.string_of_term ctxt (Logic.list_implies (asms, Logic.mk_conjunction_list props))); - fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t)) + fun cert_safe t = Thm.global_cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t)) handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; val casms = map cert_safe asms; val cprops = map cert_safe props; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/more_thm.ML Fri Mar 06 15:58:56 2015 +0100 @@ -118,7 +118,7 @@ let val thy = Thm.theory_of_cterm ct; val t = Thm.term_of ct; - in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (Thm.cterm_of thy v) | _ => I) t end; + in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (Thm.global_cterm_of thy v) | _ => I) t end; (* cterm constructors and destructors *) @@ -128,7 +128,7 @@ val thy = Thm.theory_of_cterm t; val T = Thm.typ_of_cterm t; in - Thm.apply (Thm.cterm_of thy (Const ("Pure.all", (T --> propT) --> propT))) + Thm.apply (Thm.global_cterm_of thy (Const ("Pure.all", (T --> propT) --> propT))) (Thm.lambda_name (x, t) A) end; @@ -216,7 +216,7 @@ (* type classes and sorts *) fun class_triv thy c = - Thm.of_class (Thm.ctyp_of thy (TVar ((Name.aT, 0), [c])), c); + Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; @@ -331,7 +331,7 @@ val used = fold (fn (t, u) => add_used t o add_used u) tpairs (add_used prop []); val vars = strip_vars prop; val cvars = (Name.variant_list used (map #1 vars), vars) - |> ListPair.map (fn (x, (_, T)) => Thm.cterm_of thy (Var ((x, i), T))); + |> ListPair.map (fn (x, (_, T)) => Thm.global_cterm_of thy (Var ((x, i), T))); in fold Thm.forall_elim cvars th end; in @@ -349,8 +349,8 @@ (* certify_instantiate *) fun certify_inst thy (instT, inst) = - (map (fn (v, T) => (Thm.ctyp_of thy (TVar v), Thm.ctyp_of thy T)) instT, - map (fn (v, t) => (Thm.cterm_of thy (Var v), Thm.cterm_of thy t)) inst); + (map (fn (v, T) => (Thm.global_ctyp_of thy (TVar v), Thm.global_ctyp_of thy T)) instT, + map (fn (v, t) => (Thm.global_cterm_of thy (Var v), Thm.global_cterm_of thy t)) inst); fun certify_instantiate insts th = Thm.instantiate (certify_inst (Thm.theory_of_thm th) insts) th; @@ -365,7 +365,7 @@ val fixed = fold Term.add_frees (Thm.terms_of_tpairs tpairs @ hyps) []; val frees = Term.fold_aterms (fn Free v => if member (op =) fixed v then I else insert (op =) v | _ => I) prop []; - in fold (Thm.forall_intr o Thm.cterm_of thy o Free) frees th end; + in fold (Thm.forall_intr o Thm.global_cterm_of thy o Free) frees th end; (* unvarify_global: global schematic variables *) @@ -400,7 +400,7 @@ val tfrees = rev (map TFree (Term.add_tfrees t [])); val tfrees' = map (fn a => TFree (a, [])) (Name.invent Name.context Name.aT (length tfrees)); val strip = tfrees ~~ tfrees'; - val recover = map (apply2 (Thm.ctyp_of thy o Logic.varifyT_global) o swap) strip; + val recover = map (apply2 (Thm.global_ctyp_of thy o Logic.varifyT_global) o swap) strip; val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; @@ -409,7 +409,7 @@ val _ = Sign.no_vars ctxt prop; val (strip, recover, prop') = stripped_sorts thy prop; val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; - val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of thy T, S)) strip; + val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.global_ctyp_of thy T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); @@ -426,7 +426,7 @@ fun add_def ctxt unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; - val prems = map (Thm.cterm_of thy) (Logic.strip_imp_prems prop); + val prems = map (Thm.global_cterm_of thy) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def ctxt unchecked overloaded (b, concl') thy; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/raw_simplifier.ML --- a/src/Pure/raw_simplifier.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/raw_simplifier.ML Fri Mar 06 15:58:56 2015 +0100 @@ -684,7 +684,7 @@ 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); +fun simproc_global_i thy name = mk_simproc name o map (Thm.global_cterm_of thy o Logic.varify_global); fun simproc_global thy name = simproc_global_i thy name o map (Syntax.read_term_global thy); diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/simplifier.ML Fri Mar 06 15:58:56 2015 +0100 @@ -132,7 +132,7 @@ let val lhss' = prep lthy lhss; val ctxt' = fold Variable.auto_fixes lhss' lthy; - in Variable.export_terms ctxt' lthy lhss' end |> map (Proof_Context.cterm_of lthy), + in Variable.export_terms ctxt' lthy lhss' end |> map (Thm.cterm_of lthy), proc = proc, identifier = identifier}; in diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/skip_proof.ML --- a/src/Pure/skip_proof.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/skip_proof.ML Fri Mar 06 15:58:56 2015 +0100 @@ -32,7 +32,7 @@ Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.make ("skip_proof", @{here}), I))); -fun make_thm thy prop = make_thm_cterm (Thm.cterm_of thy prop); +fun make_thm thy prop = make_thm_cterm (Thm.global_cterm_of thy prop); (* cheat_tac *) diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/subgoal.ML --- a/src/Pure/subgoal.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/subgoal.ML Fri Mar 06 15:58:56 2015 +0100 @@ -86,7 +86,7 @@ val u = Free (y, U); in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end; val (inst1, inst2) = - split_list (map (apply2 (apply2 (Proof_Context.cterm_of ctxt))) (map2 var_inst vars ys)); + split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys)); val th'' = Thm.instantiate ([], inst1) th'; in ((inst2, th''), ctxt'') end; diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/thm.ML --- a/src/Pure/thm.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/thm.ML Fri Mar 06 15:58:56 2015 +0100 @@ -23,7 +23,8 @@ (*certified types*) val theory_of_ctyp: ctyp -> theory val typ_of: ctyp -> typ - val ctyp_of: theory -> typ -> ctyp + val global_ctyp_of: theory -> typ -> ctyp + val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list (*certified terms*) val theory_of_cterm: cterm -> theory @@ -31,7 +32,8 @@ val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int - val cterm_of: theory -> term -> cterm + val global_cterm_of: theory -> term -> cterm + val cterm_of: Proof.context -> term -> cterm val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm @@ -152,13 +154,15 @@ fun theory_of_ctyp (Ctyp {thy, ...}) = thy; fun typ_of (Ctyp {T, ...}) = T; -fun ctyp_of thy raw_T = +fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {thy = thy, T = T, maxidx = maxidx, sorts = sorts} end; +val ctyp_of = global_ctyp_of o Proof_Context.theory_of; + fun dest_ctyp (Ctyp {thy, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {thy = thy, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); @@ -183,12 +187,14 @@ fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; -fun cterm_of thy tm = +fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {thy = thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; +val cterm_of = global_cterm_of o Proof_Context.theory_of; + fun merge_thys0 (Cterm {thy = thy1, ...}) (Cterm {thy = thy2, ...}) = Theory.merge (thy1, thy2); @@ -1160,7 +1166,7 @@ let val Ctyp {thy, T, ...} = cT; val c = Sign.certify_class thy raw_c; - val Cterm {t = prop, maxidx, sorts, ...} = cterm_of thy (Logic.mk_of_class (T, c)); + val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (Proofterm.OfClass (T, c)), diff -r 92d7d8e4f1bf -r 291934bac95e src/Pure/variable.ML --- a/src/Pure/variable.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Pure/variable.ML Fri Mar 06 15:58:56 2015 +0100 @@ -589,7 +589,7 @@ let val thy = Thm.theory_of_cterm goal; val ((xs, ps), ctxt') = focus_params (Thm.term_of goal) ctxt; - val ps' = map (Thm.cterm_of thy o Free) ps; + val ps' = map (Thm.global_cterm_of thy o Free) ps; val goal' = fold forall_elim_prop ps' goal; in ((xs ~~ ps', goal'), ctxt') end; diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/Code/code_preproc.ML --- a/src/Tools/Code/code_preproc.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/Code/code_preproc.ML Fri Mar 06 15:58:56 2015 +0100 @@ -153,7 +153,7 @@ fun evaluation sandwich lift_postproc eval ctxt t = let - val cert = Thm.cterm_of (Proof_Context.theory_of ctxt); + val cert = Thm.global_cterm_of (Proof_Context.theory_of ctxt); val (postproc, ct') = sandwich ctxt (cert t); in Thm.term_of ct' @@ -168,7 +168,7 @@ fun normalized_tfrees_sandwich ctxt ct = let - val cert = Thm.cterm_of (Proof_Context.theory_of ctxt); + val cert = Thm.global_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 []; @@ -185,7 +185,7 @@ fun no_variables_sandwich ctxt ct = let val thy = Proof_Context.theory_of ctxt; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val all_vars = fold_aterms (fn t as Free _ => insert (op aconvc) (cert t) | t as Var _ => insert (op aconvc) (cert t) | _ => I) (Thm.term_of ct) []; diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/Code/code_runtime.ML --- a/src/Tools/Code/code_runtime.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/Code/code_runtime.ML Fri Mar 06 15:58:56 2015 +0100 @@ -177,7 +177,7 @@ val _ = if fastype_of t <> propT then error ("Not a proposition: " ^ Syntax.string_of_term_global thy t) else (); - val iff = Thm.cterm_of thy (Term.Const (@{const_name Pure.eq}, propT --> propT --> propT)); + val iff = Thm.global_cterm_of thy (Term.Const (@{const_name Pure.eq}, propT --> propT --> propT)); val result = case partiality_as_none (evaluation truth_cookie ctxt evaluator vs_t []) of SOME Holds => true | _ => false; diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/Code/code_simp.ML --- a/src/Tools/Code/code_simp.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/Code/code_simp.ML Fri Mar 06 15:58:56 2015 +0100 @@ -83,7 +83,7 @@ THEN' conclude_tac ctxt NONE ctxt; fun dynamic_value ctxt = - snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.cterm_of (Proof_Context.theory_of ctxt); + snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.global_cterm_of (Proof_Context.theory_of ctxt); val _ = Theory.setup (Method.setup @{binding code_simp} diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/Code/code_thingol.ML Fri Mar 06 15:58:56 2015 +0100 @@ -574,7 +574,7 @@ let val raw_const = Const (c, map_type_tfree (K arity_typ') ty); val dom_length = length (fst (strip_type ty)) - val thm = Axclass.unoverload_conv thy (Thm.cterm_of thy raw_const); + val thm = Axclass.unoverload_conv thy (Thm.global_cterm_of thy raw_const); val const = (apsnd Logic.unvarifyT_global o dest_Const o snd o Logic.dest_equals o Thm.prop_of) thm; in diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/IsaPlanner/isand.ML --- a/src/Tools/IsaPlanner/isand.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/IsaPlanner/isand.ML Fri Mar 06 15:58:56 2015 +0100 @@ -91,7 +91,7 @@ fun fix_alls_cterm ctxt i th = let - val cert = Thm.cterm_of (Thm.theory_of_thm th); + val cert = Thm.global_cterm_of (Thm.theory_of_thm th); val (fixedbody, fvs) = fix_alls_term ctxt i (Thm.prop_of th); val cfvs = rev (map cert fvs); val ct_body = cert fixedbody; @@ -140,7 +140,7 @@ *) fun subgoal_thms th = let - val cert = Thm.cterm_of (Thm.theory_of_thm th); + val cert = Thm.global_cterm_of (Thm.theory_of_thm th); val t = Thm.prop_of th; diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/IsaPlanner/rw_inst.ML --- a/src/Tools/IsaPlanner/rw_inst.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/IsaPlanner/rw_inst.ML Fri Mar 06 15:58:56 2015 +0100 @@ -32,7 +32,7 @@ *) fun allify_conditions Ts th = let - val cert = Thm.cterm_of (Thm.theory_of_thm th); + val cert = Thm.global_cterm_of (Thm.theory_of_thm th); fun allify (x, T) t = Logic.all_const T $ Abs (x, T, Term.abstract_over (Free (x, T), t)); @@ -62,7 +62,7 @@ (* Note, we take abstraction in the order of last abstraction first *) fun mk_abstractedrule ctxt TsFake Ts rule = let - val cert = Thm.cterm_of (Thm.theory_of_thm rule); + val cert = Thm.global_cterm_of (Thm.theory_of_thm rule); (* now we change the names of temporary free vars that represent bound vars with binders outside the redex *) @@ -172,8 +172,8 @@ fun rw ctxt ((nonfixed_typinsts, unprepinsts), FakeTs, Ts, outerterm) rule target_thm = let val thy = Thm.theory_of_thm target_thm; - val cert = Thm.cterm_of thy; - val certT = Thm.ctyp_of thy; + val cert = Thm.global_cterm_of thy; + val certT = Thm.global_ctyp_of thy; (* fix all non-instantiated tvars *) val (fixtyinsts, othertfrees) = (* FIXME proper context!? *) diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/atomize_elim.ML --- a/src/Tools/atomize_elim.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/atomize_elim.ML Fri Mar 06 15:58:56 2015 +0100 @@ -116,7 +116,7 @@ val quantify_thesis = if is_Bound thesis then all_tac else let - val cthesis = Thm.cterm_of thy thesis + val cthesis = Thm.global_cterm_of thy thesis val rule = instantiate' [SOME (Thm.ctyp_of_cterm cthesis)] [NONE, SOME cthesis] @{thm meta_spec} in diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/coherent.ML --- a/src/Tools/coherent.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/coherent.ML Fri Mar 06 15:58:56 2015 +0100 @@ -174,8 +174,8 @@ fun thm_of_cl_prf ctxt goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = let val thy = Proof_Context.theory_of ctxt; - val cert = Thm.cterm_of thy; - val certT = Thm.ctyp_of thy; + val cert = Thm.global_cterm_of thy; + val certT = Thm.global_ctyp_of thy; val _ = cond_trace ctxt (fn () => Pretty.string_of (Pretty.big_list "asms:" (map (Display.pretty_thm ctxt) asms))); @@ -202,7 +202,7 @@ and thm_of_case_prf ctxt goal asms ((params, prf), (_, asms')) = let val thy = Proof_Context.theory_of ctxt; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val cparams = map cert params; val asms'' = map (cert o curry subst_bounds (rev params)) asms'; val (prems'', ctxt') = fold_map Thm.assume_hyps asms'' ctxt; diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/cong_tac.ML --- a/src/Tools/cong_tac.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/cong_tac.ML Fri Mar 06 15:58:56 2015 +0100 @@ -14,7 +14,7 @@ fun cong_tac ctxt cong = CSUBGOAL (fn (cgoal, i) => let - val cert = Thm.cterm_of (Thm.theory_of_cterm cgoal); + val cert = Thm.global_cterm_of (Thm.theory_of_cterm cgoal); val goal = Thm.term_of cgoal; in (case Logic.strip_assums_concl goal of diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/eqsubst.ML --- a/src/Tools/eqsubst.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/eqsubst.ML Fri Mar 06 15:58:56 2015 +0100 @@ -260,7 +260,7 @@ val tgt_term = Thm.prop_of th; val thy = Thm.theory_of_thm th; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; val cfvs = rev (map cert fvs); @@ -349,7 +349,7 @@ val tgt_term = Thm.prop_of th; val thy = Thm.theory_of_thm th; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; val cfvs = rev (map cert fvs); diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/induct.ML --- a/src/Tools/induct.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/induct.ML Fri Mar 06 15:58:56 2015 +0100 @@ -162,7 +162,7 @@ val rearrange_eqs_simproc = Simplifier.simproc_global Pure.thy "rearrange_eqs" ["Pure.all t"] - (fn ctxt => fn t => mk_swap_rrule ctxt (Thm.cterm_of (Proof_Context.theory_of ctxt) t)); + (fn ctxt => fn t => mk_swap_rrule ctxt (Thm.global_cterm_of (Proof_Context.theory_of ctxt) t)); (* rotate k premises to the left by j, skipping over first j premises *) @@ -394,7 +394,7 @@ fun prep_inst ctxt align tune (tm, ts) = let - val cert = Thm.cterm_of (Proof_Context.theory_of ctxt); + val cert = Thm.global_cterm_of (Proof_Context.theory_of ctxt); fun prep_var (x, SOME t) = let val cx = cert x; @@ -571,8 +571,8 @@ fun dest_env thy env = let - val cert = Thm.cterm_of thy; - val certT = Thm.ctyp_of thy; + val cert = Thm.global_cterm_of thy; + val certT = Thm.global_ctyp_of thy; val pairs = Vartab.dest (Envir.term_env env); val types = Vartab.dest (Envir.type_env env); val ts = map (cert o Envir.norm_term env o #2 o #2) pairs; @@ -655,7 +655,7 @@ fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) => let val thy = Proof_Context.theory_of ctxt; - val cert = Thm.cterm_of thy; + val cert = Thm.global_cterm_of thy; val v = Free (x, T); fun spec_rule prfx (xs, body) = diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/misc_legacy.ML --- a/src/Tools/misc_legacy.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/misc_legacy.ML Fri Mar 06 15:58:56 2015 +0100 @@ -161,7 +161,7 @@ fun metahyps_aux_tac ctxt tacf (prem,gno) state = let val (insts,params,hyps,concl) = metahyps_split_prem prem val maxidx = Thm.maxidx_of state - val cterm = Thm.cterm_of (Thm.theory_of_thm state) + val cterm = Thm.global_cterm_of (Thm.theory_of_thm state) val chyps = map cterm hyps val hypths = map Thm.assume chyps val subprems = map (Thm.forall_elim_vars 0) hypths @@ -265,8 +265,8 @@ let fun newName (ix,_) = (ix, gensym (string_of_indexname ix)) val alist = map newName vars fun mk_inst (v,T) = - (Thm.cterm_of thy (Var(v,T)), - Thm.cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T))) + (Thm.global_cterm_of thy (Var(v,T)), + Thm.global_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) = - (Thm.cterm_of thy (Var(v,T)), - Thm.cterm_of thy (Free(((the o AList.lookup (op =) alist) v), T))) + (Thm.global_cterm_of thy (Var(v,T)), + Thm.global_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) diff -r 92d7d8e4f1bf -r 291934bac95e src/Tools/nbe.ML --- a/src/Tools/nbe.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/Tools/nbe.ML Fri Mar 06 15:58:56 2015 +0100 @@ -78,7 +78,7 @@ val (_, triv_of_class) = Context.>>> (Context.map_theory_result (Thm.add_oracle (@{binding triv_of_class}, fn (thy, T, class) => - Thm.cterm_of thy (Logic.mk_of_class (T, class))))); + Thm.global_cterm_of thy (Logic.mk_of_class (T, class))))); in @@ -86,7 +86,7 @@ let val thy = Proof_Context.theory_of ctxt; val algebra = Sign.classes_of thy; - val certT = Thm.ctyp_of thy; + val certT = Thm.global_ctyp_of thy; val triv_classes = get_triv_classes thy; fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes; fun mk_entry (v, sort) = @@ -582,8 +582,8 @@ let val thy = Proof_Context.theory_of ctxt; val ty = Thm.typ_of_cterm lhs; - val eq = Thm.cterm_of thy (Term.Const (@{const_name Pure.eq}, ty --> ty --> propT)); - val rhs = Thm.cterm_of thy raw_rhs; + val eq = Thm.global_cterm_of thy (Term.Const (@{const_name Pure.eq}, ty --> ty --> propT)); + val rhs = Thm.global_cterm_of thy raw_rhs; in Thm.mk_binop eq lhs rhs end; val (_, raw_oracle) = Context.>>> (Context.map_theory_result diff -r 92d7d8e4f1bf -r 291934bac95e src/ZF/Tools/cartprod.ML --- a/src/ZF/Tools/cartprod.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/ZF/Tools/cartprod.ML Fri Mar 06 15:58:56 2015 +0100 @@ -106,7 +106,7 @@ fun split_rule_var ctxt (Var(v,_), Type("fun",[T1,T2]), rl) = let val T' = factors T1 ---> T2 val newt = ap_split T1 T2 (Var(v,T')) - val cterm = Thm.cterm_of (Thm.theory_of_thm rl) + val cterm = Thm.global_cterm_of (Thm.theory_of_thm rl) in remove_split ctxt (Drule.instantiate_normalize ([], [(cterm (Var(v, Ind_Syntax.iT-->T2)), cterm newt)]) rl) diff -r 92d7d8e4f1bf -r 291934bac95e src/ZF/Tools/ind_cases.ML --- a/src/ZF/Tools/ind_cases.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/ZF/Tools/ind_cases.ML Fri Mar 06 15:58:56 2015 +0100 @@ -37,7 +37,7 @@ in (case Symtab.lookup (IndCasesData.get thy) c of NONE => error ("Unknown inductive cases rule for set " ^ quote c) - | SOME f => f ctxt (Thm.cterm_of thy A)) + | SOME f => f ctxt (Thm.global_cterm_of thy A)) end; diff -r 92d7d8e4f1bf -r 291934bac95e src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML Fri Mar 06 14:01:08 2015 +0100 +++ b/src/ZF/Tools/inductive_package.ML Fri Mar 06 15:58:56 2015 +0100 @@ -494,8 +494,8 @@ The name "x.1" comes from the "RS spec" !*) val inst = case elem_frees of [_] => I - | _ => Drule.instantiate_normalize ([], [(Thm.cterm_of thy (Var(("x",1), Ind_Syntax.iT)), - Thm.cterm_of thy elem_tuple)]); + | _ => Drule.instantiate_normalize ([], [(Thm.global_cterm_of thy (Var(("x",1), Ind_Syntax.iT)), + Thm.global_cterm_of thy elem_tuple)]); (*strip quantifier and the implication*) val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));