--- a/doc-src/IsarImplementation/Thy/document/proof.tex Tue Apr 03 19:24:11 2007 +0200
+++ b/doc-src/IsarImplementation/Thy/document/proof.tex Tue Apr 03 19:24:13 2007 +0200
@@ -112,7 +112,7 @@
\indexml{Variable.declare-constraints}\verb|Variable.declare_constraints: term -> Proof.context -> Proof.context| \\
\indexml{Variable.export}\verb|Variable.export: Proof.context -> Proof.context -> thm list -> thm list| \\
\indexml{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\
- \indexml{Variable.import}\verb|Variable.import: bool -> thm list -> Proof.context ->|\isasep\isanewline%
+ \indexml{Variable.import-thms}\verb|Variable.import_thms: bool -> thm list -> Proof.context ->|\isasep\isanewline%
\verb| ((ctyp list * cterm list) * thm list) * Proof.context| \\
\indexml{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context| \\
\end{mldecls}
@@ -150,7 +150,7 @@
with \verb|Variable.polymorphic|: here the given terms are detached
from the context as far as possible.
- \item \verb|Variable.import|~\isa{open\ thms\ ctxt} invents fixed
+ \item \verb|Variable.import_thms|~\isa{open\ thms\ ctxt} invents fixed
type and term variables for the schematic ones occurring in \isa{thms}. The \isa{open} flag indicates whether the fixed names
should be accessible to the user, otherwise newly introduced names
are marked as ``internal'' (\secref{sec:names}).
--- a/doc-src/IsarImplementation/Thy/proof.thy Tue Apr 03 19:24:11 2007 +0200
+++ b/doc-src/IsarImplementation/Thy/proof.thy Tue Apr 03 19:24:13 2007 +0200
@@ -94,7 +94,7 @@
@{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
@{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
@{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
- @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
+ @{index_ML Variable.import_thms: "bool -> thm list -> Proof.context ->
((ctyp list * cterm list) * thm list) * Proof.context"} \\
@{index_ML Variable.focus: "cterm -> Proof.context -> (cterm list * cterm) * Proof.context"} \\
\end{mldecls}
@@ -133,7 +133,7 @@
with @{ML Variable.polymorphic}: here the given terms are detached
from the context as far as possible.
- \item @{ML Variable.import}~@{text "open thms ctxt"} invents fixed
+ \item @{ML Variable.import_thms}~@{text "open thms ctxt"} invents fixed
type and term variables for the schematic ones occurring in @{text
"thms"}. The @{text "open"} flag indicates whether the fixed names
should be accessible to the user, otherwise newly introduced names
--- a/src/HOL/ex/reflection.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/HOL/ex/reflection.ML Tue Apr 03 19:24:13 2007 +0200
@@ -35,7 +35,7 @@
|> fst |> strip_comb |> fst
val thy = ProofContext.theory_of ctxt
val cert = Thm.cterm_of thy
- val (((_,_),[th']), ctxt') = Variable.import true [th] ctxt
+ val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt
val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'))
fun add_fterms (t as t1 $ t2) =
if exists (fn f => could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t
--- a/src/Provers/project_rule.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Provers/project_rule.ML Tue Apr 03 19:24:13 2007 +0200
@@ -37,7 +37,7 @@
(case try imp th of
NONE => (k, th)
| SOME th' => prems (k + 1) th');
- val ((_, [rule]), ctxt') = Variable.import true [raw_rule] ctxt;
+ val ((_, [rule]), ctxt') = Variable.import_thms true [raw_rule] ctxt;
fun result i =
rule
|> proj i
@@ -57,7 +57,7 @@
(case try conj2 th of
NONE => k
| SOME th' => projs (k + 1) th');
- val ((_, [rule]), _) = Variable.import true [raw_rule] ctxt;
+ val ((_, [rule]), _) = Variable.import_thms true [raw_rule] ctxt;
in projects ctxt (1 upto projs 1 rule) raw_rule end;
end;
--- a/src/Pure/Isar/attrib.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Isar/attrib.ML Tue Apr 03 19:24:13 2007 +0200
@@ -254,7 +254,7 @@
fun standard x = no_args (Thm.rule_attribute (K Drule.standard)) x;
fun no_vars x = no_args (Thm.rule_attribute (fn ctxt => fn th =>
- let val ((_, [th']), _) = Variable.import true [th] (Context.proof_of ctxt)
+ let val ((_, [th']), _) = Variable.import_thms true [th] (Context.proof_of ctxt)
in th' end)) x;
fun eta_long x = no_args (Thm.rule_attribute (K (Drule.fconv_rule Drule.eta_long_conversion))) x;
--- a/src/Pure/Isar/element.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Isar/element.ML Tue Apr 03 19:24:13 2007 +0200
@@ -249,7 +249,7 @@
val th = MetaSimplifier.norm_hhf raw_th;
val is_elim = ObjectLogic.is_elim th;
- val ((_, [th']), ctxt') = Variable.import true [th] ctxt;
+ val ((_, [th']), ctxt') = Variable.import_thms true [th] ctxt;
val prop = Thm.prop_of th';
val (prems, concl) = Logic.strip_horn prop;
val concl_term = ObjectLogic.drop_judgment thy concl;
--- a/src/Pure/Isar/local_defs.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Isar/local_defs.ML Tue Apr 03 19:24:13 2007 +0200
@@ -124,7 +124,7 @@
(case filter_out is_trivial raw_eqs of
[] => th
| eqs =>
- let val ((_, th' :: eqs'), ctxt') = Variable.import true (th :: eqs) ctxt
+ let val ((_, th' :: eqs'), ctxt') = Variable.import_thms true (th :: eqs) ctxt
in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end);
in
--- a/src/Pure/Isar/obtain.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Isar/obtain.ML Tue Apr 03 19:24:13 2007 +0200
@@ -82,7 +82,7 @@
val _ = ObjectLogic.is_judgment thy (Thm.concl_of thm) orelse
error "Conclusion in obtained context must be object-logic judgment";
- val ((_, [thm']), ctxt') = Variable.import true [thm] fix_ctxt;
+ val ((_, [thm']), ctxt') = Variable.import_thms true [thm] fix_ctxt;
val prems = Drule.strip_imp_prems (#prop (Thm.crep_thm thm'));
in
((Drule.implies_elim_list thm' (map Thm.assume prems)
@@ -198,7 +198,7 @@
| SOME th => check_result ctxt thesis (MetaSimplifier.norm_hhf (Goal.conclude th)));
val closed_rule = Thm.forall_intr (cert (Free thesis_var)) rule;
- val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
+ val ((_, [rule']), ctxt') = Variable.import_thms false [closed_rule] ctxt;
val obtain_rule = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
val ((params, stmt), fix_ctxt) = Variable.focus (Thm.cprem_of obtain_rule 1) ctxt';
val (prems, ctxt'') =
@@ -251,7 +251,7 @@
|> Thm.forall_intr (cert (Free thesis_var))
|> Thm.instantiate (instT, []);
- val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt;
+ val ((_, rule' :: terms'), ctxt') = Variable.import_thms 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);
--- a/src/Pure/Isar/rule_cases.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Isar/rule_cases.ML Tue Apr 03 19:24:13 2007 +0200
@@ -330,7 +330,7 @@
| mutual_rule _ [th] = SOME ([0], th)
| mutual_rule ctxt (ths as th :: _) =
let
- val ((_, ths'), ctxt') = Variable.import true ths ctxt;
+ val ((_, ths'), ctxt') = Variable.import_thms true ths ctxt;
val rules as (prems, _) :: _ = map (prep_rule (get_consumes th)) ths';
val (ns, rls) = split_list (map #2 rules);
in
--- a/src/Pure/Tools/invoke.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/Tools/invoke.ML Tue Apr 03 19:24:13 2007 +0200
@@ -63,7 +63,7 @@
Proof.map_context (fn ctxt =>
let
val ([res_types, res_params, res_prems], ctxt'') =
- fold_burrow (apfst snd oo Variable.import false) results ctxt';
+ fold_burrow (apfst snd oo Variable.import_thms false) results ctxt';
val types'' = map (Logic.dest_type o Thm.term_of o Drule.dest_term) res_types;
val params'' = map (Thm.term_of o Drule.dest_term) res_params;
--- a/src/Pure/subgoal.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/subgoal.ML Tue Apr 03 19:24:13 2007 +0200
@@ -30,7 +30,7 @@
fun focus ctxt i st =
let
val ((schematics, [st']), ctxt') =
- Variable.import true [MetaSimplifier.norm_hhf_protect st] ctxt;
+ Variable.import_thms true [MetaSimplifier.norm_hhf_protect st] ctxt;
val ((params, goal), ctxt'') = Variable.focus_subgoal i st' ctxt';
val asms = Drule.strip_imp_prems goal;
val concl = Drule.strip_imp_concl goal;
--- a/src/Pure/tactic.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/tactic.ML Tue Apr 03 19:24:13 2007 +0200
@@ -118,7 +118,7 @@
fun rule_by_tactic tac rl =
let
val ctxt = Variable.thm_context rl;
- val ((_, [st]), ctxt') = Variable.import true [rl] ctxt;
+ val ((_, [st]), ctxt') = Variable.import_thms true [rl] ctxt;
in
(case Seq.pull (tac st) of
NONE => raise THM ("rule_by_tactic", 0, [rl])
--- a/src/Pure/variable.ML Tue Apr 03 19:24:11 2007 +0200
+++ b/src/Pure/variable.ML Tue Apr 03 19:24:13 2007 +0200
@@ -50,7 +50,7 @@
val import_terms: bool -> term list -> Proof.context -> term list * Proof.context
val importT: thm list -> Proof.context -> (ctyp list * thm list) * Proof.context
val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context
- val import: bool -> thm list -> Proof.context ->
+ val import_thms: bool -> thm list -> Proof.context ->
((ctyp list * cterm list) * thm list) * Proof.context
val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
@@ -392,7 +392,7 @@
val (insts, ctxt') = import_inst is_open ts ctxt;
in (Proofterm.instantiate insts prf, ctxt') end;
-fun import is_open ths ctxt =
+fun import_thms is_open ths ctxt =
let
val thy = ProofContext.theory_of ctxt;
val cert = Thm.cterm_of thy;
@@ -411,7 +411,7 @@
in exp ctxt' ctxt (f ctxt' ths') end;
val tradeT = gen_trade importT exportT;
-val trade = gen_trade (import true) export;
+val trade = gen_trade (import_thms true) export;
(* focus on outermost parameters *)