diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,468 +1,477 @@ (* Title: Pure/Isar/generic_target.ML Author: Makarius Author: Florian Haftmann, TU Muenchen Common target infrastructure. *) signature GENERIC_TARGET = sig (*auxiliary*) val export_abbrev: Proof.context -> (term -> term) -> term -> term * ((string * sort) list * (term list * term list)) val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix val check_mixfix_global: binding * bool -> mixfix -> mixfix (*background primitives*) val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val background_declaration: declaration -> local_theory -> local_theory val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> theory -> theory (*nested local theories primitives*) val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> local_theory -> local_theory val standard_declaration: (int * int -> bool) -> (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val local_interpretation: Locale.registration -> local_theory -> local_theory (*lifting target primitives to local theory operations*) val define: (((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory) -> bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val notes: (string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: (Syntax.mode -> binding * mixfix -> term -> term list * term list -> local_theory -> local_theory) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory (*theory target primitives*) val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory (*theory target operations*) val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val theory_declaration: declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory (*locale target primitives*) val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory (*locale operations*) val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val locale_dependency: string -> Locale.registration -> local_theory -> local_theory end structure Generic_Target: GENERIC_TARGET = struct (** consts **) fun export_abbrev lthy preprocess rhs = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); val rhs' = rhs |> Assumption.export_term lthy (Local_Theory.target_of lthy) |> preprocess; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); val u = fold_rev lambda term_params rhs'; val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; + val type_tfrees = Term_Subst.add_tfreesT (Term.fastype_of u) Term_Subst.TFrees.empty; val extra_tfrees = - subtract (op =) (Term.add_tfreesT (Term.fastype_of u) []) (Term.add_tfrees u []); + (u, []) |-> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I) + |> rev; val type_params = map (Logic.mk_type o TFree) extra_tfrees; in (global_rhs, (extra_tfrees, (type_params, term_params))) end; fun check_mixfix ctxt (b, extra_tfrees) mx = if null extra_tfrees then mx else (if Context_Position.is_visible ctxt then warning ("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o TFree) (sort_by #1 extra_tfrees)) ^ (if Mixfix.is_empty mx then "" else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) else (); NoSyn); fun check_mixfix_global (b, no_params) mx = if no_params orelse Mixfix.is_empty mx then mx else (warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^ Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); fun same_const (Const (c, _), Const (c', _)) = c = c' | same_const (t $ _, t' $ _) = same_const (t, t') | same_const (_, _) = false; fun const_decl phi_pred prmode ((b, mx), rhs) phi context = if phi_pred phi then let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; val same_shape = Term.aconv_untyped (rhs, rhs'); val same_stem = same_shape orelse same_const (rhs, rhs'); val const_alias = if same_shape then (case rhs' of Const (c, T) => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; in (case Type_Infer_Context.const_type ctxt c of SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE | NONE => NONE) end | _ => NONE) else NONE; in (case const_alias of SOME c => context |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |> Morphism.form (Proof_Context.generic_notation true prmode [(rhs', mx)]) | NONE => context |> Proof_Context.generic_add_abbrev Print_Mode.internal (b', Term.close_schematic_term rhs') |-> (fn (const as Const (c, _), _) => same_stem ? (Proof_Context.generic_revert_abbrev (#1 prmode) c #> same_shape ? Morphism.form (Proof_Context.generic_notation true prmode [(const, mx)])))) end else context; (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (Inttab.update_new (serial (), f)); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = Inttab.fold (fn (_, f) => f binding_const_params) interps; in lthy |> Local_Theory.background_theory (Context.theory_map interp) |> Local_Theory.map_contexts (K (Context.proof_map interp)) end; fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = let val params = type_params @ term_params; val target_params = type_params @ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); val mx' = check_mixfix_global (b, null params) mx; val (const, lthy2) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); val lhs = Term.list_comb (const, params); val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result (Thm.add_def (Proof_Context.defs_context lthy2) false false (Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) ||> foundation_interpretation (b, (const, target_params)); in ((lhs, def), lthy3) end; fun background_declaration decl lthy = let fun theory_decl context = Local_Theory.standard_form lthy (Proof_Context.init_global (Context.theory_of context)) decl context; in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; fun background_abbrev (b, global_rhs) params = Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) #>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) (** nested local theories primitives **) fun standard_facts lthy ctxt = Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); fun standard_notes pred kind facts lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd else ctxt) lthy; fun standard_declaration pred decl lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt else ctxt) lthy; fun standard_const pred prmode ((b, mx), rhs) = standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); fun standard_registration pred registration lthy = Local_Theory.map_contexts (fn level => if pred (Local_Theory.level lthy, level) then Context.proof_map (Locale.add_registration registration) else I) lthy; val local_interpretation = standard_registration (fn (n, level) => level = n - 1); (** lifting target primitives to local theory operations **) (* define *) fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) 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' []; val T = Term.fastype_of rhs; - val tfreesT = Term.add_tfreesT T (fold (Term.add_tfreesT o #2) xs []); - val extra_tfrees = rev (subtract (op =) tfreesT (Term.add_tfrees rhs [])); + val type_tfrees = + Term_Subst.TFrees.empty + |> Term_Subst.add_tfreesT T + |> fold (Term_Subst.add_tfreesT o #2) xs; + val extra_tfrees = + (rhs, []) |-> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I) + |> rev; val mx' = check_mixfix lthy (b, extra_tfrees) mx; val type_params = map (Logic.mk_type o TFree) extra_tfrees; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); val params = type_params @ term_params; val U = map Term.fastype_of params ---> T; (*foundation*) val ((lhs', global_def), lthy2) = lthy |> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); (*local definition*) val ([(lhs, (_, local_def))], lthy3) = lthy2 |> Context_Position.set_visible false |> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] ||> Context_Position.restore_visible lthy2; (*result*) val def = Thm.transitive local_def global_def |> 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; (* notes *) local fun import_export_proof ctxt (name, raw_th) = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); (*export assumes/defines*) val th = Goal.norm_result ctxt raw_th; val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; (*export fixes*) 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 Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; (*thm definition*) val result = Global_Theory.name_thm Global_Theory.official1 name th''; (*import fixes*) val (tvars, vars) = chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |>> map Logic.dest_type; val instT = fold2 (fn a => fn b => (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) tvars tfrees Term_Subst.TVars.empty; val cinstT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) instT []; val cinst = map_filter (fn (Var (xi, T), t) => SOME ((xi, Term_Subst.instantiateT instT T), Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) | _ => NONE) (vars ~~ frees); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; fun locale_target_declaration locale syntax decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_declaration locale syntax (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) (** locale operations **) fun locale_declaration locale {syntax, pervasive} decl = pervasive ? background_declaration decl #> locale_target_declaration locale syntax decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/Isar/obtain.ML b/src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML +++ b/src/Pure/Isar/obtain.ML @@ -1,419 +1,426 @@ (* Title: Pure/Isar/obtain.ML Author: Markus Wenzel, TU Muenchen Generalized existence and cases rules within Isar proof text. *) signature OBTAIN = sig val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list val obtains_attribs: ('typ, 'term) Element.obtain list -> Token.src list val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state val obtain: binding -> (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> (term * term list) list list -> (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state val obtain_cmd: binding -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> (string * string list) list list -> (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state val result: (Proof.context -> tactic) -> thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state end; structure Obtain: OBTAIN = struct (** specification elements **) (* obtain_export *) (* [x, A x] : B -------- B *) fun eliminate_term ctxt xs tm = let val vs = map (dest_Free o Thm.term_of) xs; val bads = Term.fold_aterms (fn t as Free v => if member (op =) vs v then insert (op aconv) t else I | _ => I) tm []; val _ = null bads orelse error ("Result contains obtained parameters: " ^ space_implode " " (map (Syntax.string_of_term ctxt) bads)); in tm end; fun eliminate ctxt rule xs As thm = let val _ = eliminate_term ctxt xs (Thm.full_prop_of thm); val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse error "Conclusion in obtained context must be object-logic judgment"; val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt; val prems = Drule.strip_imp_prems (Thm.cprop_of thm'); in ((Drule.implies_elim_list thm' (map Thm.assume prems) |> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As) |> Drule.forall_intr_list xs) COMP rule) |> Drule.implies_intr_list prems |> singleton (Variable.export ctxt' ctxt) end; fun obtain_export ctxt rule xs _ As = (eliminate ctxt rule xs As, eliminate_term ctxt xs); (* result declaration *) fun case_names (obtains: ('typ, 'term) Element.obtain list) = obtains |> map_index (fn (i, (b, _)) => if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b); fun obtains_attributes obtains = [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names (case_names obtains)]; fun obtains_attribs obtains = [Attrib.consumes (~ (length obtains)), Attrib.case_names (case_names obtains)]; (* obtain thesis *) fun obtain_thesis ctxt = let val ([x], ctxt') = Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt; val t = Object_Logic.fixed_judgment ctxt x; val v = dest_Free (Object_Logic.drop_judgment ctxt t); in ((v, t), ctxt') end; (* obtain clauses *) local val mk_all_external = Logic.all_constraint o Variable.default_type; fun mk_all_internal ctxt (y, z) t = let + val frees = + (t, Term_Subst.Frees.empty) + |-> Term.fold_aterms (fn Free v => Term_Subst.Frees.add (v, ()) | _ => I); val T = - (case AList.lookup (op =) (Term.add_frees t []) z of + (case Term_Subst.Frees.get_first (fn ((x, T), ()) => if x = z then SOME T else NONE) frees of SOME T => T | NONE => the_default dummyT (Variable.default_type ctxt z)); in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end; fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props = let val ((xs', vars), ctxt') = ctxt |> fold_map prep_var raw_vars |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); val xs = map (Variable.check_name o #1) vars; in Logic.list_implies (map (parse_prop ctxt') raw_props, thesis) |> fold_rev (mk_all ctxt') (xs ~~ xs') end; fun prepare_obtains prep_clause check_terms ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) = let val clauses = raw_obtains |> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props) |> check_terms ctxt; in map fst raw_obtains ~~ clauses end; val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external; val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal; in val read_obtains = prepare_obtains parse_clause Syntax.check_terms; val cert_obtains = prepare_obtains cert_clause (K I); val parse_obtains = prepare_obtains parse_clause (K I); end; (** consider: generalized elimination and cases rule **) (* consider (a) x where "A x" | (b) y where "B y" | ... \ have thesis if a [intro?]: "\x. A x \ thesis" and b [intro?]: "\y. B y \ thesis" and ... for thesis apply (insert that) *) local fun gen_consider prep_obtains raw_obtains int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains; in state |> Proof.have true NONE (K I) [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] (map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains) [((Binding.empty, atts), [(thesis, [])])] int |-> Proof.refine_insert end; in val consider = gen_consider cert_obtains; val consider_cmd = gen_consider read_obtains; end; (** obtain: augmented context based on generalized existence rule **) (* obtain (a) x where "A x" \ have thesis if a [intro?]: "\x. A x \ thesis" for thesis apply (insert that) fix x assm <> "A x" *) local fun gen_obtain prep_stmt prep_att that_binding raw_decls raw_fixes raw_prems raw_concls int state = let val _ = Proof.assert_forward_or_chain state; val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state); val ({vars, propss, binds, result_binds, ...}, params_ctxt) = prep_stmt (raw_decls @ raw_fixes) (raw_prems @ map #2 raw_concls) thesis_ctxt; val (decls, fixes) = chop (length raw_decls) vars ||> map #2; val (premss, conclss) = chop (length raw_prems) propss; val propss' = (map o map) (Logic.close_prop fixes (flat premss)) conclss; val that_prop = Logic.list_rename_params (map (#1 o #2) decls) (fold_rev (Logic.all o #2 o #2) decls (Logic.list_implies (flat propss', thesis))); val cparams = map (Thm.cterm_of params_ctxt o #2 o #2) decls; val asms = map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_concls ~~ map (map (rpair [])) propss'; fun after_qed (result_ctxt, results) state' = let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.fix (map #1 decls) |> Proof.map_context (fold (Variable.bind_term o apsnd (Logic.close_term fixes)) binds) |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms end; in state |> Proof.have true NONE after_qed [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] [((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])] [(Binding.empty_atts, [(thesis, [])])] int |-> Proof.refine_insert |> Proof.map_context (fold Variable.bind_term result_binds) end; in val obtain = gen_obtain Proof_Context.cert_stmt (K I); val obtain_cmd = gen_obtain Proof_Context.read_stmt Attrib.attribute_cmd; end; (** tactical result **) fun check_result ctxt thesis th = (case Thm.prems_of th of [prem] => if Thm.concl_of th aconv thesis andalso Logic.strip_assums_concl prem aconv thesis then th else error ("Guessed a different clause:\n" ^ Thm.string_of_thm ctxt th) | [] => error "Goal solved -- nothing guessed" | _ => error ("Guess split into several cases:\n" ^ Thm.string_of_thm ctxt th)); fun result tac facts ctxt = let val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt; val st = Goal.init (Thm.cterm_of ctxt thesis); val rule = (case SINGLE (Method.insert_tac thesis_ctxt 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 (Thm.cterm_of ctxt (Free thesis_var)) rule; val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; val obtain_rule = Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt'; val (prems, ctxt'') = Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) (Drule.strip_imp_prems stmt) fix_ctxt; in ((params, prems), ctxt'') end; (** guess: obtain based on tactical result **) (* guess x \ { fix thesis have "PROP ?guess" apply magic \ \turn goal into \thesis \ #thesis\\ apply_end magic \ \turn final \(\x. P x \ thesis) \ #thesis\ into\ \ \\#((\x. A x \ thesis) \ thesis)\ which is a finished goal state\ } fix x assm <> "A x" *) local fun unify_params vars thesis_var raw_rule ctxt = let val thy = Proof_Context.theory_of ctxt; val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); fun err msg th = error (msg ^ ":\n" ^ Thm.string_of_thm ctxt th); val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; val rule = Thm.incr_indexes (maxidx + 1) raw_rule; val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); val m = length vars; val n = length params; val _ = m <= n orelse err "More variables than parameters in obtained rule" rule; fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max) handle Type.TUNIFY => err ("Failed to unify variable " ^ string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ string_of_term (Syntax_Trans.mark_bound_abs (y, Envir.norm_type tyenv U)) ^ " in") rule; val (tyenv, _) = fold unify (map #1 vars ~~ take m params) (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule)); val norm_type = Envir.norm_type tyenv; 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 ctxt o Free) (xs @ ys'); val instT = - fold (Term.add_tvarsT o #2) params [] - |> map (fn v => (v, Thm.ctyp_of ctxt (norm_type (TVar v)))); + (params, Term_Subst.TVars.empty) |-> fold (#2 #> Term.fold_atyps (fn T => fn instT => + (case T of + TVar v => + if Term_Subst.TVars.defined instT v then instT + else Term_Subst.TVars.update (v, Thm.ctyp_of ctxt (norm_type T)) instT + | _ => instT))); val closed_rule = rule |> Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) - |> Thm.instantiate (instT, []); + |> Thm.instantiate (Term_Subst.TVars.dest 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 ctxt' (Logic.varify_global (Free thesis_var))) rule'; in ((vars', rule''), ctxt') end; fun inferred_type (binding, _, mx) ctxt = let val x = Variable.check_name binding; val ((_, T), ctxt') = Proof_Context.inferred_param x ctxt in ((x, T, mx), ctxt') end; fun polymorphic ctxt vars = let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars)) in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end; fun gen_guess prep_var raw_vars int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; val (thesis_var, thesis) = #1 (obtain_thesis ctxt); val vars = ctxt |> fold_map prep_var raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt; fun guess_context raw_rule state' = let val ((parms, rule), ctxt') = unify_params vars thesis_var raw_rule (Proof.context_of state'); val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt'; val ps = xs ~~ map (#2 o #1) parms; val ts = map Free ps; val asms = Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule)) |> map (fn asm => (Term.betapplys (fold_rev Term.abs ps asm, ts), [])); val _ = not (null asms) orelse error "Trivial result -- nothing guessed"; in state' |> 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 (Thm.cterm_of ctxt) ts)) [] [] [(Binding.empty_atts, asms)]) |> Proof.map_context (fold Variable.unbind_term Auto_Bind.no_facts) end; val goal = Var (("guess", 0), propT); val pos = Position.thread_data (); fun print_result ctxt' (k, [(s, [_, th])]) = Proof_Display.print_results int pos ctxt' (k, [(s, [th])]); val before_qed = Method.primitive_text (fn ctxt => Goal.conclude #> Raw_Simplifier.norm_hhf ctxt #> (fn th => Goal.protect 0 (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))); fun after_qed (result_ctxt, results) state' = let val [_, res] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.end_block |> guess_context (check_result ctxt thesis res) end; in state |> Proof.enter_forward |> Proof.begin_block |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |> Proof.chain_facts chain_facts |> Proof.internal_goal print_result Proof_Context.mode_schematic true "guess" (SOME before_qed) after_qed [] [] [(Binding.empty_atts, [(Logic.mk_term goal, []), (goal, [])])] |> snd |> Proof.refine_singleton (Method.primitive_text (fn _ => fn _ => Goal.init (Thm.cterm_of ctxt thesis))) end; in val guess = gen_guess Proof_Context.cert_var; val guess_cmd = gen_guess Proof_Context.read_var; end; end; diff --git a/src/Pure/Isar/specification.ML b/src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML +++ b/src/Pure/Isar/specification.ML @@ -1,478 +1,478 @@ (* Title: Pure/Isar/specification.ML Author: Makarius Derived local theory specifications --- with type-inference and toplevel polymorphism. *) signature SPECIFICATION = sig val read_props: string list -> (binding * string option * mixfix) list -> Proof.context -> term list * Proof.context val check_spec_open: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> term -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context val read_spec_open: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> string -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list val check_multi_specs: (binding * typ option * mixfix) list -> multi_specs -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val read_multi_specs: (binding * string option * mixfix) list -> multi_specs_cmd -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val axiomatization: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> (Attrib.binding * term) list -> theory -> (term list * thm list) * theory val axiomatization_cmd: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> (Attrib.binding * string) list -> theory -> (term list * thm list) * theory val axiom: Attrib.binding * term -> theory -> thm * theory val definition: (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> local_theory -> (term * (string * thm)) * local_theory val definition': (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> bool -> local_theory -> (term * (string * thm)) * local_theory val definition_cmd: (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string list -> Attrib.binding * string -> bool -> local_theory -> (term * (string * thm)) * local_theory val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term -> bool -> local_theory -> local_theory val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string -> bool -> local_theory -> local_theory val alias: binding * string -> local_theory -> local_theory val alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_alias: binding * string -> local_theory -> local_theory val type_alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val theorems: string -> (Attrib.binding * Attrib.thms) list -> (binding * typ option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorems_cmd: string -> (Attrib.binding * (Facts.ref * Token.src list) list) list -> (binding * string option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state val schematic_theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val schematic_theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state end; structure Specification: SPECIFICATION = struct (* prepare propositions *) fun read_props raw_props raw_fixes ctxt = let val (_, ctxt1) = ctxt |> Proof_Context.add_fixes_cmd raw_fixes; val props1 = map (Syntax.parse_prop ctxt1) raw_props; val (props2, ctxt2) = ctxt1 |> fold_map Variable.fix_dummy_patterns props1; val props3 = Syntax.check_props ctxt2 props2; val ctxt3 = ctxt2 |> fold Variable.declare_term props3; in (props3, ctxt3) end; (* prepare specification *) fun get_positions ctxt x = let fun get Cs (Const ("_type_constraint_", C) $ t) = get (C :: Cs) t | get Cs (Free (y, T)) = if x = y then map_filter Term_Position.decode_positionT (T :: map (Type.constraint_type ctxt) Cs) else [] | get _ (t $ u) = get [] t @ get [] u | get _ (Abs (_, _, t)) = get [] t | get _ _ = []; in get [] end; local fun prep_decls prep_var raw_vars ctxt = let val (vars, ctxt') = fold_map prep_var raw_vars ctxt; val (xs, ctxt'') = ctxt' |> Context_Position.set_visible false |> Proof_Context.add_fixes vars ||> Context_Position.restore_visible ctxt'; val _ = Context_Position.reports ctxt'' (map (Binding.pos_of o #1) vars ~~ map (Variable.markup_entity_def ctxt'' ##> Properties.remove Markup.kindN) xs); in ((vars, xs), ctxt'') end; fun close_form ctxt ys prems concl = let val xs = rev (fold (Variable.add_free_names ctxt) (prems @ [concl]) (rev ys)); val pos_props = Logic.strip_imp_concl concl :: Logic.strip_imp_prems concl @ prems; fun get_pos x = maps (get_positions ctxt x) pos_props; val _ = Context_Position.reports ctxt (maps (Syntax_Phases.reports_of_scope o get_pos) xs); in Logic.close_prop_constraint (Variable.default_type ctxt) (xs ~~ xs) prems concl end; fun dummy_frees ctxt xs tss = let val names = Variable.names_of ((fold o fold) Variable.declare_term tss ctxt) |> fold Name.declare xs; val (tss', _) = (fold_map o fold_map) Term.free_dummy_patterns tss names; in tss' end; fun prep_spec_open prep_var parse_prop raw_vars raw_params raw_prems raw_concl ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val (ys, params_ctxt) = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes; val props = map (parse_prop params_ctxt) (raw_concl :: raw_prems) |> singleton (dummy_frees params_ctxt (xs @ ys)); val concl :: prems = Syntax.check_props params_ctxt props; val spec = Logic.list_implies (prems, concl); val spec_ctxt = Variable.declare_term spec params_ctxt; fun get_pos x = maps (get_positions spec_ctxt x) props; in ((vars, xs, get_pos, spec), spec_ctxt) end; fun prep_specs prep_var parse_prop prep_att raw_vars raw_specss ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val propss0 = raw_specss |> map (fn ((_, raw_concl), raw_prems, raw_params) => let val (ys, ctxt') = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes in (ys, map (pair ctxt') (raw_concl :: raw_prems)) end); val props = burrow (grouped 10 Par_List.map_independent (uncurry parse_prop)) (map #2 propss0) |> dummy_frees vars_ctxt xs |> map2 (fn (ys, _) => fn concl :: prems => close_form vars_ctxt ys prems concl) propss0; val specs = Syntax.check_props vars_ctxt props; val specs_ctxt = vars_ctxt |> fold Variable.declare_term specs; val ps = specs_ctxt |> fold_map Proof_Context.inferred_param xs |> fst; val params = map2 (fn (b, _, mx) => fn (_, T) => ((b, T), mx)) vars ps; val name_atts: Attrib.binding list = map (fn ((name, atts), _) => (name, map (prep_att ctxt) atts)) (map #1 raw_specss); in ((params, name_atts ~~ specs), specs_ctxt) end; in val check_spec_open = prep_spec_open Proof_Context.cert_var (K I); val read_spec_open = prep_spec_open Proof_Context.read_var Syntax.parse_prop; type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list; type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list; fun check_multi_specs xs specs = prep_specs Proof_Context.cert_var (K I) (K I) xs specs; fun read_multi_specs xs specs = prep_specs Proof_Context.read_var Syntax.parse_prop Attrib.check_src xs specs; end; (* axiomatization -- within global theory *) fun gen_axioms prep_stmt prep_att raw_decls raw_fixes raw_prems raw_concls thy = let (*specification*) val ({vars, propss = [prems, concls], ...}, vars_ctxt) = Proof_Context.init_global thy |> prep_stmt (raw_decls @ raw_fixes) ((map o map) (rpair []) [raw_prems, map snd raw_concls]); val (decls, fixes) = chop (length raw_decls) vars; val frees = rev ((fold o fold) (Variable.add_frees vars_ctxt) [prems, concls] []) |> map (fn (x, T) => (x, Free (x, T))); val close = Logic.close_prop (map #2 fixes @ frees) prems; val specs = map ((apsnd o map) (prep_att vars_ctxt) o fst) raw_concls ~~ map close concls; val spec_name = Binding.conglomerate (if null decls then map (#1 o #1) specs else map (#1 o #1) decls); (*consts*) val (consts, consts_thy) = thy |> fold_map (fn ((b, _, mx), (_, t)) => Theory.specify_const ((b, Term.type_of t), mx)) decls; val subst = Term.subst_atomic (map (#2 o #2) decls ~~ consts); (*axioms*) val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), prop) => Thm.add_axiom_global (b, subst prop) #>> (fn (_, th) => ((b, atts), [([th], [])]))); (*facts*) val (facts, facts_lthy) = axioms_thy |> Named_Target.theory_init |> Spec_Rules.add spec_name Spec_Rules.Unknown consts (maps (maps #1 o #2) axioms) |> Local_Theory.notes axioms; in ((consts, map (the_single o #2) facts), Local_Theory.exit_global facts_lthy) end; val axiomatization = gen_axioms Proof_Context.cert_stmt (K I); val axiomatization_cmd = gen_axioms Proof_Context.read_stmt Attrib.check_src; fun axiom (b, ax) = axiomatization [] [] [] [(b, ax)] #>> (hd o snd); (* definition *) fun gen_def prep_spec prep_att raw_var raw_params raw_prems ((a, raw_atts), raw_spec) int lthy = let val atts = map (prep_att lthy) raw_atts; val ((vars, xs, get_pos, spec), _) = lthy |> prep_spec (the_list raw_var) raw_params raw_prems raw_spec; val (((x, T), rhs), prove) = Local_Defs.derived_def lthy get_pos {conditional = true} spec; val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val const_name = Local_Theory.full_name lthy b; val name = Thm.def_binding_optional b a; val ((lhs, (_, raw_th)), lthy2) = lthy |> Local_Theory.define_internal ((b, mx), ((Binding.suffix_name "_raw" name, []), rhs)); val th = prove lthy2 raw_th; val lthy3 = lthy2 |> Spec_Rules.add name Spec_Rules.equational [lhs] [th]; val ([(def_name, [th'])], lthy4) = lthy3 |> Local_Theory.notes [((name, atts), [([th], [])])]; val lthy5 = lthy4 |> Code.declare_default_eqns [(th', true)]; val lhs' = Morphism.term (Local_Theory.target_morphism lthy5) lhs; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy5 - (member (op =) (Term.add_frees lhs' [])) [(x, T)]; + (Term_Subst.Frees.defined (Term_Subst.add_frees lhs' Term_Subst.Frees.empty)) [(x, T)]; in ((lhs, (def_name, th')), lthy5) end; val definition' = gen_def check_spec_open (K I); fun definition xs ys As B = definition' xs ys As B false; val definition_cmd = gen_def read_spec_open Attrib.check_src; (* abbreviation *) fun gen_abbrev prep_spec mode raw_var raw_params raw_spec int lthy = let val lthy1 = lthy |> Proof_Context.set_syntax_mode mode; val ((vars, xs, get_pos, spec), _) = lthy |> Proof_Context.set_mode Proof_Context.mode_abbrev |> prep_spec (the_list raw_var) raw_params [] raw_spec; val ((x, T), rhs) = Local_Defs.abs_def (#2 (Local_Defs.cert_def lthy1 get_pos spec)); val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val lthy2 = lthy1 |> Local_Theory.abbrev mode ((b, mx), rhs) |> snd |> Proof_Context.restore_syntax_mode lthy; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy2 (K false) [(x, T)]; in lthy2 end; val abbreviation = gen_abbrev check_spec_open; val abbreviation_cmd = gen_abbrev read_spec_open; (* alias *) fun gen_alias decl check (b, arg) lthy = let val (c, reports) = check {proper = true, strict = false} lthy arg; val _ = Context_Position.reports lthy reports; in decl b c lthy end; val alias = gen_alias Local_Theory.const_alias (K (K (fn c => (c, [])))); val alias_cmd = gen_alias Local_Theory.const_alias (fn flags => fn ctxt => fn (c, pos) => apfst (#1 o dest_Const) (Proof_Context.check_const flags ctxt (c, [pos]))); val type_alias = gen_alias Local_Theory.type_alias (K (K (fn c => (c, [])))); val type_alias_cmd = gen_alias Local_Theory.type_alias (apfst (#1 o dest_Type) ooo Proof_Context.check_type_name); (* notation *) local fun gen_type_notation prep_type add mode args lthy = lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args); fun gen_notation prep_const add mode args lthy = lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* fact statements *) local fun gen_theorems prep_fact prep_att add_fixes kind raw_facts raw_fixes int lthy = let val facts = raw_facts |> map (fn ((name, atts), bs) => ((name, map (prep_att lthy) atts), bs |> map (fn (b, more_atts) => (prep_fact lthy b, map (prep_att lthy) more_atts)))); val (_, ctxt') = add_fixes raw_fixes lthy; val facts' = facts |> Attrib.partial_evaluation ctxt' |> Attrib.transform_facts (Proof_Context.export_morphism ctxt' lthy); val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts'; val _ = Proof_Display.print_results int (Position.thread_data ()) lthy' ((kind, ""), res); in (res, lthy') end; in val theorems = gen_theorems (K I) (K I) Proof_Context.add_fixes; val theorems_cmd = gen_theorems Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; end; (* complex goal statements *) local fun prep_statement prep_att prep_stmt raw_elems raw_stmt ctxt = let val (stmt, elems_ctxt) = prep_stmt raw_elems raw_stmt ctxt; val prems = Assumption.local_prems_of elems_ctxt ctxt; val stmt_ctxt = fold (fold (Proof_Context.augment o fst) o snd) stmt elems_ctxt; in (case raw_stmt of Element.Shows _ => let val stmt' = Attrib.map_specs (map prep_att) stmt in (([], prems, stmt', NONE), stmt_ctxt) end | Element.Obtains raw_obtains => let val asms_ctxt = stmt_ctxt |> fold (fn ((name, _), asm) => snd o Proof_Context.add_assms Assumption.assume_export [((name, [Context_Rules.intro_query NONE]), asm)]) stmt; val that = Assumption.local_prems_of asms_ctxt stmt_ctxt; val ([(_, that')], that_ctxt) = asms_ctxt |> Proof_Context.set_stmt true |> Proof_Context.note_thmss "" [((Binding.name Auto_Bind.thatN, []), [(that, [])])] ||> Proof_Context.restore_stmt asms_ctxt; val stmt' = [(Binding.empty_atts, [(#2 (#1 (Obtain.obtain_thesis ctxt)), [])])]; in ((Obtain.obtains_attribs raw_obtains, prems, stmt', SOME that'), that_ctxt) end) end; fun gen_theorem schematic bundle_includes prep_att prep_stmt long kind before_qed after_qed (name, raw_atts) raw_includes raw_elems raw_concl int lthy = let val _ = Local_Theory.assert lthy; val elems = raw_elems |> map (Element.map_ctxt_attrib (prep_att lthy)); val ((more_atts, prems, stmt, facts), goal_ctxt) = lthy |> bundle_includes raw_includes |> prep_statement (prep_att lthy) prep_stmt elems raw_concl; val atts = more_atts @ map (prep_att lthy) raw_atts; val pos = Position.thread_data (); fun after_qed' results goal_ctxt' = let val results' = burrow (map (Goal.norm_result lthy) o Proof_Context.export goal_ctxt' lthy) results; val (res, lthy') = if forall (Binding.is_empty_atts o fst) stmt then (map (pair "") results', lthy) else Local_Theory.notes_kind kind (map2 (fn (b, _) => fn ths => (b, [(ths, [])])) stmt results') lthy; val lthy'' = if Binding.is_empty_atts (name, atts) then (Proof_Display.print_results int pos lthy' ((kind, ""), res); lthy') else let val ([(res_name, _)], lthy'') = Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])] lthy'; val _ = Proof_Display.print_results int pos lthy' ((kind, res_name), res); in lthy'' end; in after_qed results' lthy'' end; val prems_name = if long then Auto_Bind.assmsN else Auto_Bind.thatN; in goal_ctxt |> not (null prems) ? (Proof_Context.note_thmss "" [((Binding.name prems_name, []), [(prems, [])])] #> snd) |> Proof.theorem before_qed after_qed' (map snd stmt) |> (case facts of NONE => I | SOME ths => Proof.refine_insert ths) |> tap (fn state => not schematic andalso Proof.schematic_goal state andalso error "Illegal schematic goal statement") end; in val theorem = gen_theorem false Bundle.includes (K I) Expression.cert_statement; val theorem_cmd = gen_theorem false Bundle.includes_cmd Attrib.check_src Expression.read_statement; val schematic_theorem = gen_theorem true Bundle.includes (K I) Expression.cert_statement; val schematic_theorem_cmd = gen_theorem true Bundle.includes_cmd Attrib.check_src Expression.read_statement; end; end; diff --git a/src/Pure/Isar/subgoal.ML b/src/Pure/Isar/subgoal.ML --- a/src/Pure/Isar/subgoal.ML +++ b/src/Pure/Isar/subgoal.ML @@ -1,257 +1,257 @@ (* Title: Pure/Isar/subgoal.ML Author: Makarius Author: Daniel Matichuk, NICTA/UNSW Tactical operations with explicit subgoal focus, based on canonical proof decomposition. The "visible" part of the text within the context is fixed, the remaining goal may be schematic. Isar subgoal command for proof structure within unstructured proof scripts. *) signature SUBGOAL = sig type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list} val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm val focus: Proof.context -> int -> binding list option -> thm -> focus * thm val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list -> int -> thm -> thm -> thm Seq.seq val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic val subgoal: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state val subgoal_cmd: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state end; structure Subgoal: SUBGOAL = struct (* focus *) type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list}; fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st = let val st = raw_st |> Thm.solve_constraints |> Thm.transfer' ctxt |> Raw_Simplifier.norm_hhf_protect ctxt; val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt; val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1; val (asms, concl) = if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal) else ([], goal); val text = asms @ (if do_concl then [concl] else []); val ((_, inst), ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2; val schematic_terms = Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; val schematics = (schematic_types, schematic_terms); val asms' = map (Thm.instantiate_cterm schematics) asms; val concl' = Thm.instantiate_cterm schematics concl; val (prems, context) = Assumption.add_assumes asms' ctxt3; in ({context = context, params = params, prems = prems, asms = asms', concl = concl', schematics = schematics}, Goal.init concl') end; val focus_params = gen_focus (false, false); val focus_params_fixed = gen_focus (false, true); val focus_prems = gen_focus (true, false); val focus = gen_focus (true, true); (* lift and retrofit *) (* B [?'b, ?y] ---------------- B ['b, y params] *) fun lift_import idx params th ctxt = let val ((_, [th']), ctxt') = Variable.importT [th] ctxt; val Ts = map Thm.typ_of_cterm params; val ts = map Thm.term_of params; val prop = Thm.full_prop_of th'; - val concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []; + val concl_vars = Term_Subst.add_vars (Logic.strip_imp_concl prop) Term_Subst.Vars.empty; val vars = rev (Term.add_vars prop []); val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt'; fun var_inst v y = let val ((x, i), T) = v; val (U, args) = - if member (op =) concl_vars v then (T, []) + if Term_Subst.Vars.defined concl_vars v then (T, []) else (Ts ---> T, ts); 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 (Thm.cterm_of ctxt))) (map2 var_inst vars ys)); val th'' = Thm.instantiate ([], map (apfst (Term.dest_Var o Thm.term_of)) inst1) th'; in ((inst2, th''), ctxt'') end; (* [x, A x] : B x \ C ------------------ [\x. A x \ B x] : C *) fun lift_subgoals ctxt params asms th = let fun lift ct = fold_rev (Thm.all_name ctxt) params (Drule.list_implies (asms, ct)); val unlift = fold (Thm.elim_implies o Thm.assume) asms o Drule.forall_elim_list (map #2 params) o Thm.assume; val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th)); val th' = fold (Thm.elim_implies o unlift) subgoals th; in (subgoals, th') end; fun retrofit ctxt1 ctxt0 params asms i st1 st0 = let val idx = Thm.maxidx_of st0 + 1; val ps = map #2 params; val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1; val (subgoals, st3) = lift_subgoals ctxt2 params asms st2; val result = st3 |> Goal.conclude |> Drule.implies_intr_list asms |> Drule.forall_intr_list ps |> Drule.implies_intr_list subgoals |> fold_rev (Thm.forall_intr o #1) subgoal_inst |> fold (Thm.forall_elim o #2) subgoal_inst |> Thm.adjust_maxidx_thm idx |> singleton (Variable.export ctxt2 ctxt0); in Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false} (false, result, Thm.nprems_of st1) i st0 end; (* tacticals *) fun GEN_FOCUS flags tac ctxt i st = if Thm.nprems_of st < i then Seq.empty else let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags (ctxt |> Variable.set_bound_focus true) i NONE st; in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end; val FOCUS_PARAMS = GEN_FOCUS (false, false); val FOCUS_PARAMS_FIXED = GEN_FOCUS (false, true); val FOCUS_PREMS = GEN_FOCUS (true, false); val FOCUS = GEN_FOCUS (true, true); fun SUBPROOF tac ctxt = FOCUS (Seq.map (Goal.check_finished ctxt) oo tac) ctxt; (* Isar subgoal command *) local fun param_bindings ctxt (param_suffix, raw_param_specs) st = let val _ = if Thm.no_prems st then error "No subgoals!" else (); val subgoal = #1 (Logic.dest_implies (Thm.prop_of st)); val subgoal_params = map (apfst (Name.internal o Name.clean)) (Term.strip_all_vars subgoal) |> Term.variant_frees subgoal |> map #1; val n = length subgoal_params; val m = length raw_param_specs; val _ = m <= n orelse error ("Excessive subgoal parameter specification" ^ Position.here_list (map snd (drop n raw_param_specs))); val param_specs = raw_param_specs |> map (fn (NONE, _) => NONE | (SOME x, pos) => let val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt)); val _ = Variable.check_name b; in SOME b end) |> param_suffix ? append (replicate (n - m) NONE); fun bindings (SOME x :: xs) (_ :: ys) = x :: bindings xs ys | bindings (NONE :: xs) (y :: ys) = Binding.name y :: bindings xs ys | bindings _ ys = map Binding.name ys; in bindings param_specs subgoal_params end; fun gen_subgoal prep_atts raw_result_binding raw_prems_binding param_specs state = let val _ = Proof.assert_backward state; val state1 = state |> Proof.refine_insert []; val {context = ctxt, facts = facts, goal = st} = Proof.raw_goal state1; val result_binding = apsnd (map (prep_atts ctxt)) raw_result_binding; val (prems_binding, do_prems) = (case raw_prems_binding of SOME (b, raw_atts) => ((b, map (prep_atts ctxt) raw_atts), true) | NONE => (Binding.empty_atts, false)); val (subgoal_focus, _) = (if do_prems then focus else focus_params_fixed) ctxt 1 (SOME (param_bindings ctxt param_specs st)) st; fun after_qed (ctxt'', [[result]]) = Proof.end_block #> (fn state' => let val ctxt' = Proof.context_of state'; val results' = Proof_Context.export ctxt'' ctxt' (Conjunction.elim_conjunctions result); in state' |> Proof.refine_primitive (fn _ => fn _ => retrofit ctxt'' ctxt' (#params subgoal_focus) (#asms subgoal_focus) 1 (Goal.protect 0 result) st |> Seq.hd) |> Proof.map_context (#2 o Proof_Context.note_thmss "" [(result_binding, [(results', [])])]) end) #> Proof.reset_facts #> Proof.enter_backward; in state1 |> Proof.enter_forward |> Proof.using_facts [] |> Proof.begin_block |> Proof.map_context (fn _ => #context subgoal_focus |> Proof_Context.note_thmss "" [(prems_binding, [(#prems subgoal_focus, [])])] |> #2) |> Proof.internal_goal (K (K ())) (Proof_Context.get_mode ctxt) true "subgoal" NONE after_qed [] [] [(Binding.empty_atts, [(Thm.term_of (#concl subgoal_focus), [])])] |> #2 |> Proof.using_facts facts |> pair subgoal_focus end; in val subgoal = gen_subgoal Attrib.attribute; val subgoal_cmd = gen_subgoal Attrib.attribute_cmd; end; end; val SUBPROOF = Subgoal.SUBPROOF; diff --git a/src/Pure/drule.ML b/src/Pure/drule.ML --- a/src/Pure/drule.ML +++ b/src/Pure/drule.ML @@ -1,836 +1,843 @@ (* Title: Pure/drule.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Derived rules and other operations on theorems. *) infix 0 RL RLN MRS OF COMP INCR_COMP COMP_INCR; signature BASIC_DRULE = sig val mk_implies: cterm * cterm -> cterm val list_implies: cterm list * cterm -> cterm val strip_imp_prems: cterm -> cterm list val strip_imp_concl: cterm -> cterm val cprems_of: thm -> cterm list val forall_intr_list: cterm list -> thm -> thm val forall_intr_vars: thm -> thm val forall_elim_list: cterm list -> thm -> thm val lift_all: Proof.context -> cterm -> thm -> thm val implies_elim_list: thm -> thm list -> thm val implies_intr_list: cterm list -> thm -> thm val instantiate_normalize: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate'_normalize: ctyp option list -> cterm option list -> thm -> thm val infer_instantiate_types: Proof.context -> ((indexname * typ) * cterm) list -> thm -> thm val infer_instantiate: Proof.context -> (indexname * cterm) list -> thm -> thm val infer_instantiate': Proof.context -> cterm option list -> thm -> thm val zero_var_indexes_list: thm list -> thm list val zero_var_indexes: thm -> thm val implies_intr_hyps: thm -> thm val rotate_prems: int -> thm -> thm val rearrange_prems: int list -> thm -> thm val RLN: thm list * (int * thm list) -> thm list val RL: thm list * thm list -> thm list val MRS: thm list * thm -> thm val OF: thm * thm list -> thm val COMP: thm * thm -> thm val INCR_COMP: thm * thm -> thm val COMP_INCR: thm * thm -> thm val size_of_thm: thm -> int val reflexive_thm: thm val symmetric_thm: thm val transitive_thm: thm val extensional: thm -> thm val asm_rl: thm val cut_rl: thm val revcut_rl: thm val thin_rl: thm end; signature DRULE = sig include BASIC_DRULE val outer_params: term -> (string * typ) list val generalize: Symtab.set * Symtab.set -> thm -> thm val list_comb: cterm * cterm list -> cterm val strip_comb: cterm -> cterm * cterm list val beta_conv: cterm -> cterm -> cterm val flexflex_unique: Proof.context option -> thm -> thm val export_without_context: thm -> thm val export_without_context_open: thm -> thm val store_thm: binding -> thm -> thm val store_standard_thm: binding -> thm -> thm val store_thm_open: binding -> thm -> thm val store_standard_thm_open: binding -> thm -> thm val multi_resolve: Proof.context option -> thm list -> thm -> thm Seq.seq val multi_resolves: Proof.context option -> thm list -> thm list -> thm Seq.seq val compose: thm * int * thm -> thm val equals_cong: thm val imp_cong: thm val swap_prems_eq: thm val imp_cong_rule: thm -> thm -> thm val arg_cong_rule: cterm -> thm -> thm val binop_cong_rule: cterm -> thm -> thm -> thm val fun_cong_rule: thm -> cterm -> thm val beta_eta_conversion: cterm -> thm val eta_contraction_rule: thm -> thm val norm_hhf_eq: thm val norm_hhf_eqs: thm list val is_norm_hhf: term -> bool val norm_hhf: theory -> term -> term val norm_hhf_cterm: Proof.context -> cterm -> cterm val protect: cterm -> cterm val protectI: thm val protectD: thm val protect_cong: thm val implies_intr_protected: cterm list -> thm -> thm val termI: thm val mk_term: cterm -> thm val dest_term: thm -> cterm val cterm_rule: (thm -> thm) -> cterm -> cterm val cterm_add_frees: cterm -> cterm list -> cterm list val cterm_add_vars: cterm -> cterm list -> cterm list val dummy_thm: thm val free_dummy_thm: thm val is_sort_constraint: term -> bool val sort_constraintI: thm val sort_constraint_eq: thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm val rename_bvars: (string * string) list -> thm -> thm val rename_bvars': string option list -> thm -> thm val incr_indexes: thm -> thm -> thm val incr_indexes2: thm -> thm -> thm -> thm val triv_forall_equality: thm val distinct_prems_rl: thm val equal_intr_rule: thm val equal_elim_rule1: thm val equal_elim_rule2: thm val remdups_rl: thm val abs_def: thm -> thm end; structure Drule: DRULE = struct (** some cterm->cterm operations: faster than calling cterm_of! **) (* A1\...An\B goes to [A1,...,An], where B is not an implication *) fun strip_imp_prems ct = let val (cA, cB) = Thm.dest_implies ct in cA :: strip_imp_prems cB end handle TERM _ => []; (* A1\...An\B goes to B, where B is not an implication *) fun strip_imp_concl ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => strip_imp_concl (Thm.dest_arg ct) | _ => ct); (*The premises of a theorem, as a cterm list*) val cprems_of = strip_imp_prems o Thm.cprop_of; fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t; val implies = certify Logic.implies; fun mk_implies (A, B) = Thm.apply (Thm.apply implies A) B; (*cterm version of list_implies: [A1,...,An], B goes to \A1;...;An\\B *) fun list_implies([], B) = B | list_implies(A::As, B) = mk_implies (A, list_implies(As,B)); (*cterm version of list_comb: maps (f, [t1,...,tn]) to f(t1,...,tn) *) fun list_comb (f, []) = f | list_comb (f, t::ts) = list_comb (Thm.apply f t, ts); (*cterm version of strip_comb: maps f(t1,...,tn) to (f, [t1,...,tn]) *) fun strip_comb ct = let fun stripc (p as (ct, cts)) = let val (ct1, ct2) = Thm.dest_comb ct in stripc (ct1, ct2 :: cts) end handle CTERM _ => p in stripc (ct, []) end; (*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs of the meta-equality returned by the beta_conversion rule.*) fun beta_conv x y = Thm.dest_arg (Thm.cprop_of (Thm.beta_conversion false (Thm.apply x y))); (** Standardization of rules **) (*Generalization over a list of variables*) val forall_intr_list = fold_rev Thm.forall_intr; (*Generalization over Vars -- canonical order*) fun forall_intr_vars th = fold Thm.forall_intr (Thm.add_vars th []) th; fun outer_params t = let val vs = Term.strip_all_vars t in Name.variant_list [] (map (Name.clean o #1) vs) ~~ map #2 vs end; (*lift vars wrt. outermost goal parameters -- reverses the effect of gen_all modulo higher-order unification*) fun lift_all ctxt raw_goal raw_th = let val thy = Proof_Context.theory_of ctxt; val goal = Thm.transfer_cterm thy raw_goal; val th = Thm.transfer thy raw_th; val maxidx = Thm.maxidx_of th; val ps = outer_params (Thm.term_of goal) |> 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) => ((xi, T), Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)))); + (th, Term_Subst.Vars.empty) |-> (Thm.fold_terms o Term.fold_aterms) + (fn t => fn inst => + (case t of + Var (xi, T) => + if Term_Subst.Vars.defined inst (xi, T) then inst + else + let val ct = Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)) + in Term_Subst.Vars.update ((xi, T), ct) inst end + | _ => inst)); in th - |> Thm.instantiate ([], inst) + |> Thm.instantiate ([], Term_Subst.Vars.dest inst) |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) ps end; (*direct generalization*) fun generalize names th = Thm.generalize names (Thm.maxidx_of th + 1) th; (*specialization over a list of cterms*) val forall_elim_list = fold Thm.forall_elim; (*maps A1,...,An |- B to \A1;...;An\ \ B*) val implies_intr_list = fold_rev Thm.implies_intr; (*maps \A1;...;An\ \ B and [A1,...,An] to B*) fun implies_elim_list impth ths = fold Thm.elim_implies ths impth; (*Reset Var indexes to zero, renaming to preserve distinctness*) fun zero_var_indexes_list [] = [] | zero_var_indexes_list ths = let val (instT, inst) = Term_Subst.zero_var_indexes_inst Name.context (map Thm.full_prop_of ths); val tvars = fold Thm.add_tvars ths []; fun the_tvar v = the (find_first (fn cT => v = dest_TVar (Thm.typ_of cT)) tvars); val instT' = (instT, []) |-> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => cons (v, Thm.rename_tvar b (the_tvar v))); val vars = fold (Thm.add_vars o Thm.instantiate (instT', [])) ths []; fun the_var v = the (find_first (fn ct => v = dest_Var (Thm.term_of ct)) vars); val inst' = (inst, []) |-> Term_Subst.Vars.fold (fn (v, Var (b, _)) => cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v)))); in map (Thm.adjust_maxidx_thm ~1 o Thm.instantiate (instT', inst')) ths end; val zero_var_indexes = singleton zero_var_indexes_list; (** Standard form of object-rule: no hypotheses, flexflex constraints, Frees, or outer quantifiers; all generality expressed by Vars of index 0.**) (*Discharge all hypotheses.*) fun implies_intr_hyps th = fold Thm.implies_intr (Thm.chyps_of th) th; (*Squash a theorem's flexflex constraints provided it can be done uniquely. This step can lose information.*) fun flexflex_unique opt_ctxt th = if null (Thm.tpairs_of th) then th else (case distinct Thm.eq_thm (Seq.list_of (Thm.flexflex_rule opt_ctxt th)) of [th] => th | [] => raise THM ("flexflex_unique: impossible constraints", 0, [th]) | _ => raise THM ("flexflex_unique: multiple unifiers", 0, [th])); (* old-style export without context *) val export_without_context_open = implies_intr_hyps #> Thm.forall_intr_frees #> `Thm.maxidx_of #-> (fn maxidx => Thm.forall_elim_vars (maxidx + 1) #> Thm.strip_shyps #> zero_var_indexes #> Thm.varifyT_global); val export_without_context = flexflex_unique NONE #> export_without_context_open #> Thm.close_derivation \<^here>; (*Rotates a rule's premises to the left by k*) fun rotate_prems 0 = I | rotate_prems k = Thm.permute_prems 0 k; fun with_subgoal i f = rotate_prems (i - 1) #> f #> rotate_prems (1 - i); (*Permute prems, where the i-th position in the argument list (counting from 0) gives the position within the original thm to be transferred to position i. Any remaining trailing positions are left unchanged.*) val rearrange_prems = let fun rearr new [] thm = thm | rearr new (p :: ps) thm = rearr (new + 1) (map (fn q => if new <= q andalso q < p then q + 1 else q) ps) (Thm.permute_prems (new + 1) (new - p) (Thm.permute_prems new (p - new) thm)) in rearr 0 end; (*Resolution: multiple arguments, multiple results*) local fun res opt_ctxt th i rule = (Thm.biresolution opt_ctxt false [(false, th)] i rule handle THM _ => Seq.empty) |> Seq.map Thm.solve_constraints; fun multi_res _ _ [] rule = Seq.single rule | multi_res opt_ctxt i (th :: ths) rule = Seq.maps (res opt_ctxt th i) (multi_res opt_ctxt (i + 1) ths rule); in fun multi_resolve opt_ctxt = multi_res opt_ctxt 1; fun multi_resolves opt_ctxt facts rules = Seq.maps (multi_resolve opt_ctxt facts) (Seq.of_list rules); end; (*For joining lists of rules*) fun thas RLN (i, thbs) = let val resolve = Thm.biresolution NONE false (map (pair false) thas) i fun resb thb = Seq.list_of (resolve thb) handle THM _ => [] in maps resb thbs |> map Thm.solve_constraints end; fun thas RL thbs = thas RLN (1, thbs); (*Isar-style multi-resolution*) fun bottom_rl OF rls = (case Seq.chop 2 (multi_resolve NONE rls bottom_rl) of ([th], _) => Thm.solve_constraints th | ([], _) => raise THM ("OF: no unifiers", 0, bottom_rl :: rls) | _ => raise THM ("OF: multiple unifiers", 0, bottom_rl :: rls)); (*Resolve a list of rules against bottom_rl from right to left; makes proof trees*) fun rls MRS bottom_rl = bottom_rl OF rls; (*compose Q and \...,Qi,Q(i+1),...\ \ R to \...,Q(i+1),...\ \ R with no lifting or renaming! Q may contain \ or meta-quantifiers ALWAYS deletes premise i *) fun compose (tha, i, thb) = Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, tha, 0) i thb |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("compose: unique result expected", i, [tha, thb])); (** theorem equality **) (*Useful "distance" function for BEST_FIRST*) val size_of_thm = size_of_term o Thm.full_prop_of; (*** Meta-Rewriting Rules ***) val read_prop = certify o Simple_Syntax.read_prop; fun store_thm name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm (name, th))); fun store_thm_open name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm_open (name, th))); fun store_standard_thm name th = store_thm name (export_without_context th); fun store_standard_thm_open name th = store_thm_open name (export_without_context_open th); val reflexive_thm = let val cx = certify (Var(("x",0),TVar(("'a",0),[]))) in store_standard_thm_open (Binding.make ("reflexive", \<^here>)) (Thm.reflexive cx) end; val symmetric_thm = let val xy = read_prop "x::'a \ y::'a"; val thm = Thm.implies_intr xy (Thm.symmetric (Thm.assume xy)); in store_standard_thm_open (Binding.make ("symmetric", \<^here>)) thm end; val transitive_thm = let val xy = read_prop "x::'a \ y::'a"; val yz = read_prop "y::'a \ z::'a"; val xythm = Thm.assume xy; val yzthm = Thm.assume yz; val thm = Thm.implies_intr yz (Thm.transitive xythm yzthm); in store_standard_thm_open (Binding.make ("transitive", \<^here>)) thm end; fun extensional eq = let val eq' = Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (Thm.cprop_of eq)))) eq in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end; val equals_cong = store_standard_thm_open (Binding.make ("equals_cong", \<^here>)) (Thm.reflexive (read_prop "x::'a \ y::'a")); val imp_cong = let val ABC = read_prop "A \ B::prop \ C::prop" val AB = read_prop "A \ B" val AC = read_prop "A \ C" val A = read_prop "A" in store_standard_thm_open (Binding.make ("imp_cong", \<^here>)) (Thm.implies_intr ABC (Thm.equal_intr (Thm.implies_intr AB (Thm.implies_intr A (Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.implies_elim (Thm.assume AB) (Thm.assume A))))) (Thm.implies_intr AC (Thm.implies_intr A (Thm.equal_elim (Thm.symmetric (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))) (Thm.implies_elim (Thm.assume AC) (Thm.assume A))))))) end; val swap_prems_eq = let val ABC = read_prop "A \ B \ C" val BAC = read_prop "B \ A \ C" val A = read_prop "A" val B = read_prop "B" in store_standard_thm_open (Binding.make ("swap_prems_eq", \<^here>)) (Thm.equal_intr (Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A (Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B))))) (Thm.implies_intr BAC (Thm.implies_intr A (Thm.implies_intr B (Thm.implies_elim (Thm.implies_elim (Thm.assume BAC) (Thm.assume B)) (Thm.assume A)))))) end; val imp_cong_rule = Thm.combination o Thm.combination (Thm.reflexive implies); fun arg_cong_rule ct th = Thm.combination (Thm.reflexive ct) th; (*AP_TERM in LCF/HOL*) fun fun_cong_rule th ct = Thm.combination th (Thm.reflexive ct); (*AP_THM in LCF/HOL*) fun binop_cong_rule ct th1 th2 = Thm.combination (arg_cong_rule ct th1) th2; fun beta_eta_conversion ct = let val thm = Thm.beta_conversion true ct in Thm.transitive thm (Thm.eta_conversion (Thm.rhs_of thm)) end; (*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*) fun eta_contraction_rule th = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of th)) th; (* abs_def *) (* f ?x1 ... ?xn \ u -------------------- f \ \x1 ... xn. u *) local fun contract_lhs th = Thm.transitive (Thm.symmetric (beta_eta_conversion (#1 (Thm.dest_equals (Thm.cprop_of th))))) th; fun var_args ct = (case try Thm.dest_comb ct of SOME (f, arg) => (case Thm.term_of arg of Var ((x, _), _) => update (eq_snd (op aconvc)) (x, arg) (var_args f) | _ => []) | NONE => []); in fun abs_def th = let val th' = contract_lhs th; val args = var_args (Thm.lhs_of th'); in contract_lhs (fold (uncurry Thm.abstract_rule) args th') end; end; (*** Some useful meta-theorems ***) (*The rule V/V, obtains assumption solving for eresolve_tac*) val asm_rl = store_standard_thm_open (Binding.make ("asm_rl", \<^here>)) (Thm.trivial (read_prop "?psi")); (*Meta-level cut rule: \V \ W; V\ \ W *) val cut_rl = store_standard_thm_open (Binding.make ("cut_rl", \<^here>)) (Thm.trivial (read_prop "?psi \ ?theta")); (*Generalized elim rule for one conclusion; cut_rl with reversed premises: \PROP V; PROP V \ PROP W\ \ PROP W *) val revcut_rl = let val V = read_prop "V"; val VW = read_prop "V \ W"; in store_standard_thm_open (Binding.make ("revcut_rl", \<^here>)) (Thm.implies_intr V (Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V)))) end; (*for deleting an unwanted assumption*) val thin_rl = let val V = read_prop "V"; val W = read_prop "W"; val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W)); in store_standard_thm_open (Binding.make ("thin_rl", \<^here>)) thm end; (* (\x. PROP ?V) \ PROP ?V Allows removal of redundant parameters*) val triv_forall_equality = let val V = read_prop "V"; val QV = read_prop "\x::'a. V"; val x = certify (Free ("x", Term.aT [])); in store_standard_thm_open (Binding.make ("triv_forall_equality", \<^here>)) (Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV))) (Thm.implies_intr V (Thm.forall_intr x (Thm.assume V)))) end; (* (PROP ?Phi \ PROP ?Phi \ PROP ?Psi) \ (PROP ?Phi \ PROP ?Psi) *) val distinct_prems_rl = let val AAB = read_prop "Phi \ Phi \ Psi"; val A = read_prop "Phi"; in store_standard_thm_open (Binding.make ("distinct_prems_rl", \<^here>)) (implies_intr_list [AAB, A] (implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A])) end; (* \PROP ?phi \ PROP ?psi; PROP ?psi \ PROP ?phi\ \ PROP ?phi \ PROP ?psi Introduction rule for \ as a meta-theorem. *) val equal_intr_rule = let val PQ = read_prop "phi \ psi"; val QP = read_prop "psi \ phi"; in store_standard_thm_open (Binding.make ("equal_intr_rule", \<^here>)) (Thm.implies_intr PQ (Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP)))) end; (* PROP ?phi \ PROP ?psi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule1 = let val eq = read_prop "phi::prop \ psi::prop"; val P = read_prop "phi"; in store_standard_thm_open (Binding.make ("equal_elim_rule1", \<^here>)) (Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P]) end; (* PROP ?psi \ PROP ?phi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule2 = store_standard_thm_open (Binding.make ("equal_elim_rule2", \<^here>)) (symmetric_thm RS equal_elim_rule1); (* PROP ?phi \ PROP ?phi \ PROP ?psi \ PROP ?psi *) val remdups_rl = let val P = read_prop "phi"; val Q = read_prop "psi"; val thm = implies_intr_list [P, P, Q] (Thm.assume Q); in store_standard_thm_open (Binding.make ("remdups_rl", \<^here>)) thm end; (** embedded terms and types **) local val A = certify (Free ("A", propT)); val axiom = Thm.unvarify_axiom (Context.the_global_context ()); val prop_def = axiom "Pure.prop_def"; val term_def = axiom "Pure.term_def"; val sort_constraint_def = axiom "Pure.sort_constraint_def"; val C = Thm.lhs_of sort_constraint_def; val T = Thm.dest_arg C; val CA = mk_implies (C, A); in (* protect *) val protect = Thm.apply (certify Logic.protectC); val protectI = store_standard_thm (Binding.concealed (Binding.make ("protectI", \<^here>))) (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A)); val protectD = store_standard_thm (Binding.concealed (Binding.make ("protectD", \<^here>))) (Thm.equal_elim prop_def (Thm.assume (protect A))); val protect_cong = store_standard_thm_open (Binding.make ("protect_cong", \<^here>)) (Thm.reflexive (protect A)); fun implies_intr_protected asms th = let val asms' = map protect asms in implies_elim_list (implies_intr_list asms th) (map (fn asm' => Thm.assume asm' RS protectD) asms') |> implies_intr_list asms' end; (* term *) val termI = store_standard_thm (Binding.concealed (Binding.make ("termI", \<^here>))) (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A))); fun mk_term ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; in Thm.instantiate ([((("'a", 0), []), cT)], [((("x", 0), T), ct)]) termI end; fun dest_term th = let val cprop = strip_imp_concl (Thm.cprop_of th) in if can Logic.dest_term (Thm.term_of cprop) then Thm.dest_arg cprop else raise THM ("dest_term", 0, [th]) end; fun cterm_rule f = dest_term o f o mk_term; val cterm_add_frees = Thm.add_frees o mk_term; val cterm_add_vars = Thm.add_vars o mk_term; val dummy_thm = mk_term (certify Term.dummy_prop); val free_dummy_thm = Thm.tag_free_dummy dummy_thm; (* sort_constraint *) fun is_sort_constraint (Const ("Pure.sort_constraint", _) $ Const ("Pure.type", _)) = true | is_sort_constraint _ = false; val sort_constraintI = store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", \<^here>))) (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T)); val sort_constraint_eq = store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", \<^here>))) (Thm.equal_intr (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify_global (Context.the_global_context ()) sort_constraintI))) (implies_intr_list [A, C] (Thm.assume A))); end; (* HHF normalization *) (* (PROP ?phi \ (\x. PROP ?psi x)) \ (\x. PROP ?phi \ PROP ?psi x) *) val norm_hhf_eq = let val aT = TFree ("'a", []); val x = Free ("x", aT); val phi = Free ("phi", propT); val psi = Free ("psi", aT --> propT); val cx = certify x; val cphi = certify phi; val lhs = certify (Logic.mk_implies (phi, Logic.all x (psi $ x))); val rhs = certify (Logic.all x (Logic.mk_implies (phi, psi $ x))); in Thm.equal_intr (Thm.implies_elim (Thm.assume lhs) (Thm.assume cphi) |> Thm.forall_elim cx |> Thm.implies_intr cphi |> Thm.forall_intr cx |> Thm.implies_intr lhs) (Thm.implies_elim (Thm.assume rhs |> Thm.forall_elim cx) (Thm.assume cphi) |> Thm.forall_intr cx |> Thm.implies_intr cphi |> Thm.implies_intr rhs) |> store_standard_thm_open (Binding.make ("norm_hhf_eq", \<^here>)) end; val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq); val norm_hhf_eqs = [norm_hhf_eq, sort_constraint_eq]; fun is_norm_hhf (Const ("Pure.sort_constraint", _)) = false | is_norm_hhf (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false | is_norm_hhf (Abs _ $ _) = false | is_norm_hhf (t $ u) = is_norm_hhf t andalso is_norm_hhf u | is_norm_hhf (Abs (_, _, t)) = is_norm_hhf t | is_norm_hhf _ = true; fun norm_hhf thy t = if is_norm_hhf t then t else Pattern.rewrite_term thy [norm_hhf_prop] [] t; fun norm_hhf_cterm ctxt raw_ct = let val thy = Proof_Context.theory_of ctxt; val ct = Thm.transfer_cterm thy raw_ct; val t = Thm.term_of ct; in if is_norm_hhf t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end; (* var indexes *) fun incr_indexes th = Thm.incr_indexes (Thm.maxidx_of th + 1); fun incr_indexes2 th1 th2 = Thm.incr_indexes (Int.max (Thm.maxidx_of th1, Thm.maxidx_of th2) + 1); local (*compose Q and \Q1,Q2,...,Qk\ \ R to \Q2,...,Qk\ \ R getting unique result*) fun comp incremented th1 th2 = Thm.bicompose NONE {flatten = true, match = false, incremented = incremented} (false, th1, 0) 1 th2 |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("COMP", 1, [th1, th2])); in fun th1 COMP th2 = comp false th1 th2; fun th1 INCR_COMP th2 = comp true (incr_indexes th2 th1) th2; fun th1 COMP_INCR th2 = comp true th1 (incr_indexes th1 th2); end; fun comp_no_flatten (th, n) i rule = (case distinct Thm.eq_thm (Seq.list_of (Thm.bicompose NONE {flatten = false, match = false, incremented = true} (false, th, n) i (incr_indexes th rule))) of [th'] => Thm.solve_constraints th' | [] => raise THM ("comp_no_flatten", i, [th, rule]) | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule])); (** variations on Thm.instantiate **) fun instantiate_normalize instpair th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate instpair th COMP_INCR asm_rl); fun instantiate'_normalize Ts ts th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate' Ts ts th COMP_INCR asm_rl); (*instantiation with type-inference for variables*) fun infer_instantiate_types _ [] th = th | infer_instantiate_types ctxt args raw_th = let val thy = Proof_Context.theory_of ctxt; val th = Thm.transfer thy raw_th; fun infer ((xi, T), cu) (tyenv, maxidx) = let val _ = Thm.ctyp_of ctxt T; val _ = Thm.transfer_cterm thy cu; val U = Thm.typ_of_cterm cu; val maxidx' = maxidx |> Integer.max (#2 xi) |> Term.maxidx_typ T |> Integer.max (Thm.maxidx_of_cterm cu); val (tyenv', maxidx'') = Sign.typ_unify thy (T, U) (tyenv, maxidx') handle Type.TUNIFY => let val t = Var (xi, T); val u = Thm.term_of cu; in raise THM ("infer_instantiate_types: type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv T) ^ " of variable " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) t) ^ "\ncannot be unified with type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv U) ^ " of term " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) u), 0, [th]) end; in (tyenv', maxidx'') end; val (tyenv, _) = fold infer args (Vartab.empty, 0); val instT = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), Thm.ctyp_of ctxt (Envir.norm_type tyenv T))) tyenv []; val inst = args |> map (fn ((xi, T), cu) => ((xi, Envir.norm_type tyenv T), Thm.instantiate_cterm (instT, []) (Thm.transfer_cterm thy cu))); in instantiate_normalize (instT, inst) th end handle CTERM (msg, _) => raise THM (msg, 0, [raw_th]) | TERM (msg, _) => raise THM (msg, 0, [raw_th]) | TYPE (msg, _, _) => raise THM (msg, 0, [raw_th]); fun infer_instantiate _ [] th = th | infer_instantiate ctxt args th = let val vars = Term.add_vars (Thm.full_prop_of th) []; val dups = duplicates (eq_fst op =) vars; val _ = null dups orelse raise THM ("infer_instantiate: inconsistent types for variables " ^ commas_quote (map (Syntax.string_of_term (Config.put show_types true ctxt) o Var) dups), 0, [th]); val args' = args |> map_filter (fn (xi, cu) => AList.lookup (op =) vars xi |> Option.map (fn T => ((xi, T), cu))); in infer_instantiate_types ctxt args' th end; fun infer_instantiate' ctxt args th = let val vars = rev (Term.add_vars (Thm.full_prop_of th) []); val args' = zip_options vars args handle ListPair.UnequalLengths => raise THM ("infer_instantiate': more instantiations than variables in thm", 0, [th]); in infer_instantiate_types ctxt args' th end; (** renaming of bound variables **) (* replace bound variables x_i in thm by y_i *) (* where vs = [(x_1, y_1), ..., (x_n, y_n)] *) fun rename_bvars [] thm = thm | rename_bvars vs thm = let fun rename (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; (* renaming in left-to-right order *) fun rename_bvars' xs thm = let fun rename [] t = ([], t) | rename (x' :: xs) (Abs (x, T, t)) = let val (xs', t') = rename xs t in (xs', Abs (the_default x x', T, t')) end | rename xs (t $ u) = let val (xs', t') = rename xs t; val (xs'', u') = rename xs' u; in (xs'', t' $ u') end | rename xs a = (xs, a); in (case rename xs (Thm.prop_of thm) of ([], prop') => Thm.renamed_prop prop' thm | _ => error "More names than abstractions in theorem") end; end; structure Basic_Drule: BASIC_DRULE = Drule; open Basic_Drule; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,746 +1,763 @@ (* Title: Pure/more_thm.ML Author: Makarius Further operations on type ctyp/cterm/thm, outside the inference kernel. *) infix aconvc; signature BASIC_THM = sig include BASIC_THM val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T structure Ctermtab: TABLE structure Thmtab: TABLE val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM structure Ctermtab: TABLE structure Thmtab: TABLE val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool val add_tvars: thm -> ctyp list -> ctyp list val add_frees: thm -> cterm list -> cterm list val add_vars: thm -> cterm list -> cterm list val dest_funT: ctyp -> ctyp * ctyp val strip_type: ctyp -> ctyp list * ctyp val all_name: Proof.context -> string * cterm -> cterm -> cterm val all: Proof.context -> cterm -> cterm -> cterm val mk_binop: cterm -> cterm -> cterm -> cterm val dest_binop: cterm -> cterm * cterm val dest_implies: cterm -> cterm * cterm val dest_equals: cterm -> cterm * cterm val dest_equals_lhs: cterm -> cterm val dest_equals_rhs: cterm -> cterm val lhs_of: thm -> cterm val rhs_of: thm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val thm_ord: thm ord val cterm_cache: (cterm -> 'a) -> cterm -> 'a val thm_cache: (thm -> 'a) -> thm -> 'a val is_reflexive: thm -> bool val eq_thm: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val eq_thm_strict: thm * thm -> bool val equiv_thm: theory -> thm * thm -> bool val class_triv: theory -> class -> thm val of_sort: ctyp * sort -> thm list val is_dummy: thm -> bool val add_thm: thm -> thm list -> thm list val del_thm: thm -> thm list -> thm list val merge_thms: thm list * thm list -> thm list val item_net: thm Item_Net.T val item_net_intro: thm Item_Net.T val item_net_elim: thm Item_Net.T val declare_hyps: cterm -> Proof.context -> Proof.context val assume_hyps: cterm -> Proof.context -> thm * Proof.context val unchecked_hyps: Proof.context -> Proof.context val restore_hyps: Proof.context -> Proof.context -> Proof.context val undeclared_hyps: Context.generic -> thm -> term list val check_hyps: Context.generic -> thm -> thm val declare_term_sorts: term -> Proof.context -> Proof.context val extra_shyps': Proof.context -> thm -> sort list val check_shyps: Proof.context -> thm -> thm val weaken_sorts': Proof.context -> cterm -> cterm val elim_implies: thm -> thm -> thm val forall_intr_name: string * cterm -> thm -> thm val forall_elim_var: int -> thm -> thm val forall_elim_vars: int -> thm -> thm val instantiate_frees: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> thm -> thm val instantiate': ctyp option list -> cterm option list -> thm -> thm val forall_intr_frees: thm -> thm val unvarify_global: theory -> thm -> thm val unvarify_axiom: theory -> string -> thm val rename_params_rule: string list * int -> thm -> thm val rename_boundvars: term -> term -> thm -> thm val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory val add_axiom_global: binding * term -> theory -> (string * thm) * theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory type attribute = Context.generic * thm -> Context.generic option * thm option type binding = binding * attribute list val tag_rule: string * string -> thm -> thm val untag_rule: string -> thm -> thm val is_free_dummy: thm -> bool val tag_free_dummy: thm -> thm val def_name: string -> string val def_name_optional: string -> string -> string val def_binding: Binding.binding -> Binding.binding val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding val make_def_binding: bool -> Binding.binding -> Binding.binding val has_name_hint: thm -> bool val get_name_hint: thm -> string val put_name_hint: string -> thm -> thm val theoremK: string val legacy_get_kind: thm -> string val kind_rule: string -> thm -> thm val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic val theory_attributes: attribute list -> thm -> theory -> thm * theory val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context val no_attributes: 'a -> 'a * 'b list val simple_fact: 'a -> ('a * 'b list) list val tag: string * string -> attribute val untag: string -> attribute val kind: string -> attribute val register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val expose_theory: theory -> unit val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val pretty_thm_item: Proof.context -> thm -> Pretty.T val pretty_thm_global: theory -> thm -> Pretty.T val string_of_thm: Proof.context -> thm -> string val string_of_thm_global: theory -> thm -> string end; structure Thm: THM = struct (** basic operations **) (* collecting ctyps and cterms *) val eq_ctyp = op = o apply2 Thm.typ_of; val op aconvc = op aconv o apply2 Thm.term_of; val add_tvars = Thm.fold_atomic_ctyps (fn a => is_TVar (Thm.typ_of a) ? insert eq_ctyp a); val add_frees = Thm.fold_atomic_cterms (fn a => is_Free (Thm.term_of a) ? insert (op aconvc) a); val add_vars = Thm.fold_atomic_cterms (fn a => is_Var (Thm.term_of a) ? insert (op aconvc) a); (* ctyp operations *) fun dest_funT cT = (case Thm.typ_of cT of Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end | T => raise TYPE ("dest_funT", [T], [])); (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) fun strip_type cT = (case Thm.typ_of cT of Type ("fun", _) => let val (cT1, cT2) = dest_funT cT; val (cTs, cT') = strip_type cT2 in (cT1 :: cTs, cT') end | _ => ([], cT)); (* cterm operations *) fun all_name ctxt (x, t) A = let val T = Thm.typ_of_cterm t; val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT)); in Thm.apply all_const (Thm.lambda_name (x, t) A) end; fun all ctxt t A = all_name ctxt ("", t) A; fun mk_binop c a b = Thm.apply (Thm.apply c a) b; fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); fun dest_implies ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_implies", [Thm.term_of ct])); fun dest_equals ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_equals", [Thm.term_of ct])); fun dest_equals_lhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); fun dest_equals_rhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); val lhs_of = dest_equals_lhs o Thm.cprop_of; val rhs_of = dest_equals_rhs o Thm.cprop_of; (* certified term order *) val fast_term_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of; val term_ord = Term_Ord.term_ord o apply2 Thm.term_of; (* thm order: ignores theory context! *) val thm_ord = Term_Ord.fast_term_ord o apply2 Thm.prop_of ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 Thm.tpairs_of ||| list_ord Term_Ord.fast_term_ord o apply2 Thm.hyps_of ||| list_ord Term_Ord.sort_ord o apply2 Thm.shyps_of; (* tables and caches *) structure Ctermtab = Table(type key = cterm val ord = fast_term_ord); structure Thmtab = Table(type key = thm val ord = thm_ord); fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o thm_ord; val eq_thm_prop = op aconv o apply2 Thm.full_prop_of; fun eq_thm_strict ths = eq_thm ths andalso Context.eq_thy_id (apply2 Thm.theory_id ths) andalso op = (apply2 Thm.maxidx_of ths) andalso op = (apply2 Thm.get_tags ths); (* pattern equivalence *) fun equiv_thm thy ths = Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths); (* type classes and sorts *) fun class_triv thy 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; (* misc operations *) fun is_dummy thm = (case try Logic.dest_term (Thm.concl_of thm) of NONE => false | SOME t => Term.is_dummy_pattern (Term.head_of t)); (* collections of theorems in canonical order *) val add_thm = update eq_thm_prop; val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; val item_net = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val item_net_intro = Item_Net.init eq_thm_prop (single o Thm.concl_of); val item_net_elim = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); (** declared hyps and sort hyps **) structure Hyps = Proof_Data ( type T = {checked_hyps: bool, hyps: Termtab.set, shyps: sort Ord_List.T}; fun init _ : T = {checked_hyps = true, hyps = Termtab.empty, shyps = []}; ); fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} => let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end); (* hyps *) fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct; val hyps' = Termtab.update (Thm.term_of ct, ()) hyps; in (checked_hyps, hyps', shyps) end); fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); fun restore_hyps ctxt = map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps)); fun undeclared_hyps context th = Thm.hyps_of th |> filter_out (case context of Context.Theory _ => K false | Context.Proof ctxt => (case Hyps.get ctxt of {checked_hyps = false, ...} => K true | {hyps, ...} => Termtab.defined hyps)); fun check_hyps context th = (case undeclared_hyps context th of [] => th | undeclared => error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" (map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared)))); (* shyps *) fun declare_term_sorts t = map_hyps (fn (checked_hyps, hyps, shyps) => (checked_hyps, hyps, Sorts.insert_term t shyps)); fun extra_shyps' ctxt th = Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); fun check_shyps ctxt raw_th = let val th = Thm.strip_shyps raw_th; val extra_shyps = extra_shyps' ctxt th; in if null extra_shyps then th else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) end; val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; (** basic derived rules **) (*Elimination of implication A A \ B ------------ B *) fun elim_implies thA thAB = Thm.implies_elim thAB thA; (* forall_intr_name *) fun forall_intr_name (a, x) th = let val th' = Thm.forall_intr x th; val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); in Thm.renamed_prop prop' th' end; (* forall_elim_var(s) *) local fun dest_all ct = (case Thm.term_of ct of Const ("Pure.all", _) $ Abs (a, _, _) => let val (x, ct') = Thm.dest_abs NONE (Thm.dest_arg ct) in SOME ((a, Thm.ctyp_of_cterm x), ct') end | _ => NONE); fun dest_all_list ct = (case dest_all ct of NONE => [] | SOME (v, ct') => v :: dest_all_list ct'); fun forall_elim_vars_list vars i th = let val used = (Thm.fold_terms o Term.fold_aterms) (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I) th []; val vars' = (Name.variant_list used (map #1 vars), vars) |> ListPair.map (fn (x, (_, T)) => Thm.var ((x, i), T)); in fold Thm.forall_elim vars' th end; in fun forall_elim_vars i th = forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th; fun forall_elim_var i th = let val vars = (case dest_all (Thm.cprop_of th) of SOME (v, _) => [v] | NONE => raise THM ("forall_elim_var", i, [th])); in forall_elim_vars_list vars i th end; end; (* instantiate frees *) fun instantiate_frees ([], []) th = th | instantiate_frees (instT, inst) th = let val idx = Thm.maxidx_of th + 1; fun index ((a, A), b) = (((a, idx), A), b); val insts = (map index instT, map index inst); val tfrees = fold (Symtab.insert_set o #1 o #1) instT Symtab.empty; val frees = fold (Symtab.insert_set o #1 o #1) inst Symtab.empty; val hyps = Thm.chyps_of th; val inst_cterm = Thm.generalize_cterm (tfrees, frees) idx #> Thm.instantiate_cterm insts; in th |> fold_rev Thm.implies_intr hyps |> Thm.generalize (tfrees, frees) idx |> Thm.instantiate insts |> fold (elim_implies o Thm.assume o inst_cterm) hyps end; (* instantiate by left-to-right occurrence of variables *) fun instantiate' cTs cts thm = let fun err msg = raise TYPE ("instantiate': " ^ msg, map_filter (Option.map Thm.typ_of) cTs, map_filter (Option.map Thm.term_of) cts); fun zip_vars xs ys = zip_options xs ys handle ListPair.UnequalLengths => err "more instantiations than variables in thm"; val thm' = Thm.instantiate ((zip_vars (rev (Thm.fold_terms Term.add_tvars thm [])) cTs), []) thm; val thm'' = Thm.instantiate ([], zip_vars (rev (Thm.fold_terms Term.add_vars thm' [])) cts) thm'; in thm'' end; (* forall_intr_frees: generalization over all suitable Free variables *) fun forall_intr_frees th = let val fixed = - fold Term.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th) @ Thm.hyps_of th) []; + Term_Subst.Frees.empty + |> fold Term_Subst.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) + |> fold Term_Subst.add_frees (Thm.hyps_of th); val frees = Thm.fold_atomic_cterms (fn a => (case Thm.term_of a of - Free v => not (member (op =) fixed v) ? insert (op aconvc) a + Free v => not (Term_Subst.Frees.defined fixed v) ? insert (op aconvc) a | _ => I)) th []; in fold Thm.forall_intr frees th end; (* unvarify_global: global schematic variables *) fun unvarify_global thy th = let val prop = Thm.full_prop_of th; val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); - val instT = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S))); - val instantiateT = Term_Subst.instantiateT (Term_Subst.TVars.table instT); - val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => - let val T' = instantiateT T - in (((a, i), T'), Thm.global_cterm_of thy (Free ((a, T')))) end); - in Thm.instantiate (map (apsnd (Thm.global_ctyp_of thy)) instT, inst) th end; + val cert = Thm.global_cterm_of thy; + val certT = Thm.global_ctyp_of thy; + + val instT = + (prop, Term_Subst.TVars.empty) |-> (Term.fold_types o Term.fold_atyps) + (fn T => fn instT => + (case T of + TVar (v as ((a, _), S)) => + if Term_Subst.TVars.defined instT v then instT + else Term_Subst.TVars.update (v, TFree (a, S)) instT + | _ => instT)); + val cinstT = Term_Subst.TVars.map (K certT) instT; + val cinst = + (prop, Term_Subst.Vars.empty) |-> Term.fold_aterms + (fn t => fn inst => + (case t of + Var ((x, i), T) => + let val T' = Term_Subst.instantiateT instT T + in Term_Subst.Vars.update (((x, i), T'), cert (Free ((x, T')))) inst end + | _ => inst)); + in Thm.instantiate (Term_Subst.TVars.dest cinstT, Term_Subst.Vars.dest cinst) th end; fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; (* user renaming of parameters in a subgoal *) (*The names, if distinct, are used for the innermost parameters of subgoal i; preceding parameters may be renamed to make all parameters distinct.*) fun rename_params_rule (names, i) st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val params = map #1 (Logic.strip_params Bi); val short = length params - length names; val names' = if short < 0 then error "More names than parameters in subgoal!" else Name.variant_list names (take short params) @ names; val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val Bi' = Logic.list_rename_params names' Bi; in (case duplicates (op =) names of a :: _ => (warning ("Can't rename. Bound variables not distinct: " ^ a); st) | [] => (case inter (op =) names free_names of a :: _ => (warning ("Can't rename. Bound/Free variable clash: " ^ a); st) | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) end; (* preservation of bound variable names *) fun rename_boundvars pat obj th = (case Term.rename_abs pat obj (Thm.prop_of th) of NONE => th | SOME prop' => Thm.renamed_prop prop' th); (** specification primitives **) (* rules *) fun stripped_sorts thy t = let val tfrees = rev (Term.add_tfrees t []); val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); val recover = map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S)))) tfrees' tfrees; val strip = map (apply2 TFree) (tfrees ~~ tfrees'); val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; fun add_axiom ctxt (b, prop) thy = let 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 ctxt T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold elim_implies of_sorts; in ((axm_name, thm), thy') end; fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def context unchecked overloaded (b, concl') thy; val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold_rev Thm.implies_intr prems; in ((axm_name, thm), thy') end; fun add_def_global unchecked overloaded arg thy = add_def (Defs.global_context thy) unchecked overloaded arg thy; (** theorem tags **) (* add / delete tags *) fun tag_rule tg = Thm.map_tags (insert (op =) tg); fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); (* free dummy thm -- for abstract closure *) val free_dummyN = "free_dummy"; fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; val tag_free_dummy = tag_rule (free_dummyN, ""); (* def_name *) fun def_name c = c ^ "_def"; fun def_name_optional c "" = def_name c | def_name_optional _ name = name; val def_binding = Binding.map_name def_name #> Binding.reset_pos; fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); (* theorem kinds *) val theoremK = "theorem"; fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; (** attributes **) (*attributes subsume any kind of rules or context modifiers*) type attribute = Context.generic * thm -> Context.generic option * thm option; type binding = binding * attribute list; fun rule_attribute ths f (x, th) = (NONE, (case find_first is_free_dummy (th :: ths) of SOME th' => SOME th' | NONE => SOME (f x th))); fun declaration_attribute f (x, th) = (if is_free_dummy th then NONE else SOME (f th x), NONE); fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; fun apply_attribute (att: attribute) th x = let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) in (the_default th th', the_default x x') end; fun attribute_declaration att th x = #2 (apply_attribute att th x); fun apply_attributes mk dest = let fun app [] th x = (th, x) | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; in app end; val theory_attributes = apply_attributes Context.Theory Context.the_theory; val proof_attributes = apply_attributes Context.Proof Context.the_proof; fun no_attributes x = (x, []); fun simple_fact x = [(x, [])]; fun tag tg = rule_attribute [] (K (tag_rule tg)); fun untag s = rule_attribute [] (K (untag_rule s)); fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k)); (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun reset_proofs thy = if Inttab.is_empty (Proofs.get thy) then NONE else SOME (Proofs.put Inttab.empty thy); val _ = Theory.setup (Theory.at_begin reset_proofs); fun register_proofs ths thy = let val entry = (serial (), Lazy.map_finished (map Thm.trim_context) ths) in (Proofs.map o Inttab.update) entry thy end; fun force_proofs thy = Proofs.get thy |> Inttab.dest |> maps (map (Thm.transfer thy) o Lazy.force o #2); val consolidate_theory = Thm.consolidate o force_proofs; fun expose_theory thy = if Proofterm.export_enabled () then Thm.expose_proofs thy (force_proofs thy) else (); (** print theorems **) (* options *) val show_consts = Config.declare_option_bool ("show_consts", \<^here>); val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false); val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false); (* pretty_thm etc. *) fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th = let val show_tags = Config.get ctxt show_tags; val show_hyps = Config.get ctxt show_hyps; val th = raw_th |> perhaps (try (Thm.transfer' ctxt)) |> perhaps (try Thm.strip_shyps); val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; val extra_shyps = extra_shyps' ctxt th; val tags = Thm.get_tags th; val tpairs = Thm.tpairs_of th; val q = if quote then Pretty.quote else I; val prt_term = q o Syntax.pretty_term ctxt; val hlen = length extra_shyps + length hyps + length tpairs; val hsymbs = if hlen = 0 then [] else if show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" (map (q o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ map (Syntax.pretty_sort ctxt) extra_shyps)] else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")]; val tsymbs = if null tags orelse not show_tags then [] else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true}; fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; fun pretty_thm_global thy = pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false}; val string_of_thm = Pretty.string_of oo pretty_thm; val string_of_thm_global = Pretty.string_of oo pretty_thm_global; open Thm; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/primitive_defs.ML b/src/Pure/primitive_defs.ML --- a/src/Pure/primitive_defs.ML +++ b/src/Pure/primitive_defs.ML @@ -1,85 +1,85 @@ (* Title: Pure/primitive_defs.ML Author: Makarius Primitive definition forms. *) signature PRIMITIVE_DEFS = sig val dest_def: Proof.context -> {check_head: term -> bool, check_free_lhs: string -> bool, check_free_rhs: string -> bool, check_tfree: string -> bool} -> term -> (term * term) * term list * term val abs_def: term -> term * term end; structure Primitive_Defs: PRIMITIVE_DEFS = struct fun term_kind (Const _) = "existing constant " | term_kind (Free _) = "free variable " | term_kind (Bound _) = "bound variable " | term_kind _ = ""; (*c x \ t[x] to \x. c x \ t[x]*) fun dest_def ctxt {check_head, check_free_lhs, check_free_rhs, check_tfree} eq = let fun err msg = raise TERM (msg, [eq]); val eq_vars = Term.strip_all_vars eq; val eq_body = Term.strip_all_body eq; val display_terms = commas_quote o map (Syntax.string_of_term ctxt o Syntax_Trans.bound_vars eq_vars); val display_types = commas_quote o map (Syntax.string_of_typ ctxt); val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (\)"; val lhs = Envir.beta_eta_contract raw_lhs; val (head, args) = Term.strip_comb lhs; - val head_tfrees = Term.add_tfrees head []; + val head_tfrees = Term_Subst.add_tfrees head Term_Subst.TFrees.empty; fun check_arg (Bound _) = true | check_arg (Free (x, _)) = check_free_lhs x | check_arg (Const ("Pure.type", Type ("itself", [TFree _]))) = true | check_arg _ = false; fun close_arg (Bound _) t = t | close_arg x t = Logic.all x t; val lhs_bads = filter_out check_arg args; val lhs_dups = duplicates (op aconv) args; val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => if check_free_rhs x orelse member (op aconv) args v then I else insert (op aconv) v | _ => I) rhs []; val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => - if check_tfree a orelse member (op =) head_tfrees (a, S) then I + if check_tfree a orelse Term_Subst.TFrees.defined head_tfrees (a, S) then I else insert (op =) v | _ => I)) rhs []; in if not (check_head head) then err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head]) else if not (null lhs_bads) then err ("Bad arguments on lhs: " ^ display_terms lhs_bads) else if not (null lhs_dups) then err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups) else if not (null rhs_extras) then err ("Extra variables on rhs: " ^ display_terms rhs_extras) else if not (null rhs_extrasT) then err ("Extra type variables on rhs: " ^ display_types rhs_extrasT) else if exists_subterm (fn t => t aconv head) rhs then err "Entity to be defined occurs on rhs" else ((lhs, rhs), args, fold_rev close_arg args (Logic.list_all (eq_vars, (Logic.mk_equals (lhs, rhs))))) end; (*\x. c x \ t[x] to c \ \x. t[x]*) fun abs_def eq = let val body = Term.strip_all_body eq; val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq)); val (lhs, rhs) = Logic.dest_equals (Term.subst_bounds (vars, body)); val (lhs', args) = Term.strip_comb lhs; val rhs' = fold_rev (absfree o dest_Free) args rhs; in (lhs', rhs') end; end; diff --git a/src/Pure/theory.ML b/src/Pure/theory.ML --- a/src/Pure/theory.ML +++ b/src/Pure/theory.ML @@ -1,342 +1,342 @@ (* Title: Pure/theory.ML Author: Lawrence C Paulson and Markus Wenzel Logical theory content: axioms, definitions, and begin/end wrappers. *) signature THEORY = sig val parents_of: theory -> theory list val ancestors_of: theory -> theory list val nodes_of: theory -> theory list val setup: (theory -> theory) -> unit val local_setup: (Proof.context -> Proof.context) -> unit val install_pure: theory -> unit val get_pure: unit -> theory val get_pure_bootstrap: unit -> theory val get_markup: theory -> Markup.T val check: {long: bool} -> Proof.context -> string * Position.T -> theory val axiom_table: theory -> term Name_Space.table val axiom_space: theory -> Name_Space.T val all_axioms_of: theory -> (string * term) list val defs_of: theory -> Defs.T val at_begin: (theory -> theory option) -> theory -> theory val at_end: (theory -> theory option) -> theory -> theory val join_theory: theory list -> theory val begin_theory: string * Position.T -> theory list -> theory val end_theory: theory -> theory val add_axiom: Proof.context -> binding * term -> theory -> theory val const_dep: theory -> string * typ -> Defs.entry val type_dep: string * typ list -> Defs.entry val add_deps: Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory val add_deps_global: string -> Defs.entry -> Defs.entry list -> theory -> theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> theory val specify_const: (binding * typ) * mixfix -> theory -> term * theory val check_overloading: Proof.context -> bool -> string * typ -> unit end structure Theory: THEORY = struct (** theory context operations **) val parents_of = Context.parents_of; val ancestors_of = Context.ancestors_of; fun nodes_of thy = thy :: ancestors_of thy; fun setup f = Context.>> (Context.map_theory f); fun local_setup f = Context.>> (Context.map_proof f); (* implicit theory Pure *) val pure: theory Single_Assignment.var = Single_Assignment.var "pure"; fun install_pure thy = Single_Assignment.assign pure thy; fun get_pure () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => raise Fail "Theory Pure not present"); fun get_pure_bootstrap () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => Context.the_global_context ()); (** datatype thy **) type wrapper = (theory -> theory option) * stamp; fun apply_wrappers (wrappers: wrapper list) = perhaps (perhaps_loop (perhaps_apply (map fst wrappers))); datatype thy = Thy of {pos: Position.T, id: serial, axioms: term Name_Space.table, defs: Defs.T, wrappers: wrapper list * wrapper list}; fun make_thy (pos, id, axioms, defs, wrappers) = Thy {pos = pos, id = id, axioms = axioms, defs = defs, wrappers = wrappers}; structure Thy = Theory_Data' ( type T = thy; val extend = I; val empty = make_thy (Position.none, 0, Name_Space.empty_table Markup.axiomN, Defs.empty, ([], [])); fun merge old_thys (thy1, thy2) = let val Thy {pos, id, axioms = axioms1, defs = defs1, wrappers = (bgs1, ens1)} = thy1; val Thy {pos = _, id = _, axioms = axioms2, defs = defs2, wrappers = (bgs2, ens2)} = thy2; val axioms' = Name_Space.merge_tables (axioms1, axioms2); val defs' = Defs.merge (Defs.global_context (fst old_thys)) (defs1, defs2); val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2); val ens' = Library.merge (eq_snd op =) (ens1, ens2); in make_thy (pos, id, axioms', defs', (bgs', ens')) end; ); fun rep_theory thy = Thy.get thy |> (fn Thy args => args); fun map_thy f = Thy.map (fn (Thy {pos, id, axioms, defs, wrappers}) => make_thy (f (pos, id, axioms, defs, wrappers))); fun map_axioms f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, f axioms, defs, wrappers)); fun map_defs f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, f defs, wrappers)); fun map_wrappers f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, defs, f wrappers)); (* entity markup *) fun theory_markup def name id pos = if id = 0 then Markup.empty else Position.make_entity_markup def id Markup.theoryN (name, pos); fun init_markup (name, pos) thy = let val id = serial (); val _ = Context_Position.reports_global thy [(pos, theory_markup true name id pos)]; in map_thy (fn (_, _, axioms, defs, wrappers) => (pos, id, axioms, defs, wrappers)) thy end; fun get_markup thy = let val {pos, id, ...} = rep_theory thy in theory_markup false (Context.theory_long_name thy) id pos end; fun check long ctxt (name, pos) = let val thy = Proof_Context.theory_of ctxt; val thy' = Context.get_theory long thy name handle ERROR msg => let val completion_report = Completion.make_report (name, pos) (fn completed => map (Context.theory_name' long) (ancestors_of thy) |> filter (completed o Long_Name.base_name) |> sort_strings |> map (fn a => (a, (Markup.theoryN, a)))); in error (msg ^ Position.here pos ^ completion_report) end; val _ = Context_Position.report ctxt pos (get_markup thy'); in thy' end; (* basic operations *) val axiom_table = #axioms o rep_theory; val axiom_space = Name_Space.space_of_table o axiom_table; val all_axioms_of = Name_Space.dest_table o axiom_table; val defs_of = #defs o rep_theory; (* join theory *) fun join_theory [] = raise List.Empty | join_theory [thy] = thy | join_theory thys = foldl1 Context.join_thys thys; (* begin/end theory *) val begin_wrappers = rev o #1 o #wrappers o rep_theory; val end_wrappers = rev o #2 o #wrappers o rep_theory; fun at_begin f = map_wrappers (apfst (cons (f, stamp ()))); fun at_end f = map_wrappers (apsnd (cons (f, stamp ()))); fun begin_theory (name, pos) imports = if name = Context.PureN then (case imports of [thy] => init_markup (name, pos) thy | _ => error "Bad bootstrapping of theory Pure") else let val thy = Context.begin_thy name imports; val wrappers = begin_wrappers thy; in thy |> init_markup (name, pos) |> Sign.init_naming |> Sign.local_path |> apply_wrappers wrappers |> tap (Syntax.force_syntax o Sign.syn_of) end; fun end_theory thy = thy |> apply_wrappers (end_wrappers thy) |> Sign.change_check |> Context.finish_thy; (** primitive specifications **) (* raw axioms *) fun cert_axm ctxt (b, raw_tm) = let val thy = Proof_Context.theory_of ctxt; val t = Sign.cert_prop thy raw_tm handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg; val _ = Term.no_dummy_patterns t handle TERM (msg, _) => error msg; val bad_sorts = rev ((fold_types o fold_atyps_sorts) (fn (_, []) => I | (T, _) => insert (op =) T) t []); val _ = null bad_sorts orelse error ("Illegal sort constraints in primitive specification: " ^ commas (map (Syntax.string_of_typ (Config.put show_sorts true ctxt)) bad_sorts)); in (b, Sign.no_vars ctxt t) end handle ERROR msg => cat_error msg ("The error(s) above occurred in axiom " ^ Binding.print b); fun add_axiom ctxt raw_axm thy = thy |> map_axioms (fn axioms => let val axm = apsnd Logic.varify_global (cert_axm ctxt raw_axm); val context = ctxt |> Sign.inherit_naming thy |> Context_Position.set_visible_generic false; val (_, axioms') = Name_Space.define context true axm axioms; in axioms' end); (* dependencies *) fun const_dep thy (c, T) = ((Defs.Const, c), Sign.const_typargs thy (c, T)); fun type_dep (c, args) = ((Defs.Type, c), args); fun dependencies (context as (ctxt, _)) unchecked def description lhs rhs = let fun prep (item, args) = (case fold Term.add_tvarsT args [] of [] => (item, map Logic.varifyT_global args) | vs => raise TYPE ("Illegal schematic type variable(s)", map TVar vs, [])); - val lhs_vars = fold Term.add_tfreesT (snd lhs) []; + val lhs_vars = fold Term_Subst.add_tfreesT (snd lhs) Term_Subst.TFrees.empty; val rhs_extras = fold (fn (_, args) => args |> (fold o Term.fold_atyps) (fn TFree v => - if member (op =) lhs_vars v then I else insert (op =) v)) rhs []; + if Term_Subst.TFrees.defined lhs_vars v then I else insert (op =) v)) rhs []; val _ = if null rhs_extras then () else error ("Specification depends on extra type variables: " ^ commas_quote (map (Syntax.string_of_typ ctxt o TFree) rhs_extras) ^ "\nThe error(s) above occurred in " ^ quote description); in Defs.define context unchecked def description (prep lhs) (map prep rhs) end; fun cert_entry thy ((Defs.Const, c), args) = Sign.cert_term thy (Const (c, Sign.const_instance thy (c, args))) |> dest_Const |> const_dep thy | cert_entry thy ((Defs.Type, c), args) = Sign.certify_typ thy (Type (c, args)) |> dest_Type |> type_dep; fun add_deps context a raw_lhs raw_rhs thy = let val (lhs as ((_, lhs_name), _)) :: rhs = map (cert_entry thy) (raw_lhs :: raw_rhs); val description = if a = "" then lhs_name ^ " axiom" else a; in thy |> map_defs (dependencies context false NONE description lhs rhs) end; fun add_deps_global a x y thy = add_deps (Defs.global_context thy) a x y thy; fun specify_const decl thy = let val (t, thy') = Sign.declare_const_global decl thy; in (t, add_deps_global "" (const_dep thy' (dest_Const t)) [] thy') end; (* overloading *) fun check_overloading ctxt overloaded (c, T) = let val thy = Proof_Context.theory_of ctxt; val declT = Sign.the_const_constraint thy c handle TYPE (msg, _, _) => error msg; val T' = Logic.varifyT_global T; fun message sorts txt = [Pretty.block [Pretty.str "Specification of constant ", Pretty.str c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Config.put show_sorts sorts ctxt) T)], Pretty.str txt] |> Pretty.chunks |> Pretty.string_of; in if Sign.typ_instance thy (declT, T') then () else if Type.raw_instance (declT, T') then error (message true "imposes additional sort constraints on the constant declaration") else if overloaded then () else error (message false "is strictly less general than the declared type (overloading required)") end; (* definitional axioms *) local fun check_def (context as (ctxt, _)) thy unchecked overloaded (b, tm) defs = let val name = Sign.full_name thy b; val ((lhs, rhs), _, _) = Primitive_Defs.dest_def ctxt {check_head = Term.is_Const, check_free_lhs = K true, check_free_rhs = K false, check_tfree = K false} tm handle TERM (msg, _) => error msg; val lhs_const = Term.dest_Const (Term.head_of lhs); val rhs_consts = fold_aterms (fn Const const => insert (op =) (const_dep thy const) | _ => I) rhs []; val rhs_types = (fold_types o fold_subtypes) (fn Type t => insert (op =) (type_dep t) | _ => I) rhs []; val rhs_deps = rhs_consts @ rhs_types; val _ = check_overloading ctxt overloaded lhs_const; in defs |> dependencies context unchecked (SOME name) name (const_dep thy lhs_const) rhs_deps end handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block [Pretty.str ("The error(s) above occurred in definition " ^ Binding.print b ^ ":"), Pretty.fbrk, Pretty.quote (Syntax.pretty_term ctxt tm)])); in fun add_def (context as (ctxt, _)) unchecked overloaded raw_axm thy = let val axm = cert_axm ctxt raw_axm in thy |> map_defs (check_def context thy unchecked overloaded axm) |> add_axiom ctxt axm end; end; end; diff --git a/src/Pure/variable.ML b/src/Pure/variable.ML --- a/src/Pure/variable.ML +++ b/src/Pure/variable.ML @@ -1,748 +1,748 @@ (* Title: Pure/variable.ML Author: Makarius Fixed type/term variables and polymorphic term abbreviations. *) signature VARIABLE = sig val names_of: Proof.context -> Name.context val binds_of: Proof.context -> (typ * term) Vartab.table val maxidx_of: Proof.context -> int val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table val is_declared: Proof.context -> string -> bool val check_name: binding -> string val default_type: Proof.context -> string -> typ option val def_type: Proof.context -> bool -> indexname -> typ option val def_sort: Proof.context -> indexname -> sort option val declare_maxidx: int -> Proof.context -> Proof.context val declare_names: term -> Proof.context -> Proof.context val declare_constraints: term -> Proof.context -> Proof.context val declare_internal: term -> Proof.context -> Proof.context val declare_term: term -> Proof.context -> Proof.context val declare_typ: typ -> Proof.context -> Proof.context val declare_prf: Proofterm.proof -> Proof.context -> Proof.context val declare_thm: thm -> Proof.context -> Proof.context val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list val bind_term: indexname * term -> Proof.context -> Proof.context val unbind_term: indexname -> Proof.context -> Proof.context val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context val expand_binds: Proof.context -> term -> term val lookup_const: Proof.context -> string -> string option val is_const: Proof.context -> string -> bool val declare_const: string * string -> Proof.context -> Proof.context val next_bound: string * typ -> Proof.context -> term * Proof.context val revert_bounds: Proof.context -> term -> term val is_body: Proof.context -> bool val set_body: bool -> Proof.context -> Proof.context val restore_body: Proof.context -> Proof.context -> Proof.context val improper_fixes: Proof.context -> Proof.context val restore_proper_fixes: Proof.context -> Proof.context -> Proof.context val is_improper: Proof.context -> string -> bool val is_fixed: Proof.context -> string -> bool val is_newly_fixed: Proof.context -> Proof.context -> string -> bool val fixed_ord: Proof.context -> string ord val intern_fixed: Proof.context -> string -> string val lookup_fixed: Proof.context -> string -> string option val revert_fixed: Proof.context -> string -> string val markup_fixed: Proof.context -> string -> Markup.T val markup: Proof.context -> string -> Markup.T val markup_entity_def: Proof.context -> string -> Markup.T val dest_fixes: Proof.context -> (string * string) list val add_fixed_names: Proof.context -> term -> string list -> string list val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_newly_fixed: Proof.context -> Proof.context -> term -> (string * typ) list -> (string * typ) list val add_free_names: Proof.context -> term -> string list -> string list val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context val add_fixes: string list -> Proof.context -> string list * Proof.context val add_fixes_direct: string list -> Proof.context -> Proof.context val add_fixes_implicit: term -> Proof.context -> Proof.context val fix_dummy_patterns: term -> Proof.context -> term * Proof.context val variant_fixes: string list -> Proof.context -> string list * Proof.context val gen_all: Proof.context -> thm -> thm val export_terms: Proof.context -> Proof.context -> term list -> term list val exportT_terms: Proof.context -> Proof.context -> term list -> term list val exportT: Proof.context -> Proof.context -> thm list -> thm list val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof val export: Proof.context -> Proof.context -> thm list -> thm list val export_morphism: Proof.context -> Proof.context -> morphism val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context val importT_inst: term list -> Proof.context -> typ Term_Subst.TVars.table * Proof.context val import_inst: bool -> term list -> Proof.context -> (typ Term_Subst.TVars.table * term Term_Subst.Vars.table) * Proof.context val importT_terms: term list -> Proof.context -> term list * Proof.context val import_terms: bool -> term list -> Proof.context -> term list * Proof.context val importT: thm list -> Proof.context -> (((indexname * sort) * 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 -> ((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) * Proof.context val import_vars: Proof.context -> thm -> thm 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 val is_bound_focus: Proof.context -> bool val set_bound_focus: bool -> Proof.context -> Proof.context val restore_bound_focus: Proof.context -> Proof.context -> Proof.context val focus_params: binding list option -> term -> Proof.context -> (string list * (string * typ) list) * Proof.context val focus: binding list option -> term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context val focus_cterm: binding list option -> cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val focus_subgoal: binding list option -> int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val warn_extra_tfrees: Proof.context -> Proof.context -> unit val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list val polymorphic: Proof.context -> term list -> term list end; structure Variable: VARIABLE = struct (** local context data **) type fixes = (string * bool) Name_Space.table; val empty_fixes: fixes = Name_Space.empty_table Markup.fixedN; datatype data = Data of {names: Name.context, (*type/term variable names*) consts: string Symtab.table, (*consts within the local scope*) bounds: int * ((string * typ) * string) list, (*next index, internal name, type, external name*) fixes: fixes, (*term fixes -- global name space, intern ~> extern*) binds: (typ * term) Vartab.table, (*term bindings*) type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) maxidx: int, (*maximum var index*) constraints: typ Vartab.table * (*type constraints*) sort Vartab.table}; (*default sorts*) fun make_data (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) = Data {names = names, consts = consts, bounds = bounds, fixes = fixes, binds = binds, type_occs = type_occs, maxidx = maxidx, constraints = constraints}; val empty_data = make_data (Name.context, Symtab.empty, (0, []), empty_fixes, Vartab.empty, Symtab.empty, ~1, (Vartab.empty, Vartab.empty)); structure Data = Proof_Data ( type T = data; fun init _ = empty_data; ); fun map_data f = Data.map (fn Data {names, consts, bounds, fixes, binds, type_occs, maxidx, constraints} => make_data (f (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints))); fun map_names f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (f names, consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_consts f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, f consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_bounds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, f bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_fixes f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, f fixes, binds, type_occs, maxidx, constraints)); fun map_binds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, f binds, type_occs, maxidx, constraints)); fun map_type_occs f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, f type_occs, maxidx, constraints)); fun map_maxidx f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, f maxidx, constraints)); fun map_constraints f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, maxidx, f constraints)); fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); val names_of = #names o rep_data; val fixes_of = #fixes o rep_data; val fixes_space = Name_Space.space_of_table o fixes_of; val binds_of = #binds o rep_data; val type_occs_of = #type_occs o rep_data; val maxidx_of = #maxidx o rep_data; val constraints_of = #constraints o rep_data; val is_declared = Name.is_declared o names_of; val check_name = Name_Space.base_name o tap Binding.check; (** declarations **) (* default sorts and types *) fun default_type ctxt x = Vartab.lookup (#1 (constraints_of ctxt)) (x, ~1); fun def_type ctxt pattern xi = let val {binds, constraints = (types, _), ...} = rep_data ctxt in (case Vartab.lookup types xi of NONE => if pattern then NONE else Vartab.lookup binds xi |> Option.map (Type.mark_polymorphic o #1) | some => some) end; val def_sort = Vartab.lookup o #2 o constraints_of; (* maxidx *) val declare_maxidx = map_maxidx o Integer.max; (* names *) fun declare_type_names t = map_names (fold_types (fold_atyps Term.declare_typ_names) t) #> map_maxidx (fold_types Term.maxidx_typ t); fun declare_names t = declare_type_names t #> map_names (fold_aterms Term.declare_term_frees t) #> map_maxidx (Term.maxidx_term t); (* type occurrences *) fun decl_type_occsT T = fold_atyps (fn TFree (a, _) => Symtab.default (a, []) | _ => I) T; val decl_type_occs = fold_term_types (fn Free (x, _) => fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) | _ => decl_type_occsT); val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; val declare_type_occs = map_type_occs o decl_type_occs; (* constraints *) fun constrain_tvar (xi, raw_S) = let val S = #2 (Term_Position.decode_positionS raw_S) in if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S) end; fun declare_constraints t = map_constraints (fn (types, sorts) => let val types' = fold_aterms (fn Free (x, T) => Vartab.update ((x, ~1), T) | Var v => Vartab.update v | _ => I) t types; val sorts' = (fold_types o fold_atyps) (fn TFree (x, S) => constrain_tvar ((x, ~1), S) | TVar v => constrain_tvar v | _ => I) t sorts; in (types', sorts') end) #> declare_type_occsT t #> declare_type_names t; (* common declarations *) fun declare_internal t = declare_names t #> declare_type_occs t #> Thm.declare_term_sorts t; fun declare_term t = declare_internal t #> declare_constraints t; val declare_typ = declare_term o Logic.mk_type; val declare_prf = Proofterm.fold_proof_terms_types declare_internal (declare_internal o Logic.mk_type); val declare_thm = Thm.fold_terms declare_internal; (* renaming term/type frees *) fun variant_frees ctxt ts frees = let val names = names_of (fold declare_names ts ctxt); val xs = fst (fold_map Name.variant (map #1 frees) names); in xs ~~ map snd frees end; (** term bindings **) fun bind_term ((x, i), t) = let val u = Term.close_schematic_term t; val U = Term.fastype_of u; in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; val unbind_term = map_binds o Vartab.delete_safe; fun maybe_bind_term (xi, SOME t) = bind_term (xi, t) | maybe_bind_term (xi, NONE) = unbind_term xi; fun expand_binds ctxt = let val binds = binds_of ctxt; val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; in Envir.beta_norm o Envir.expand_term get end; (** consts **) val lookup_const = Symtab.lookup o #consts o rep_data; val is_const = is_some oo lookup_const; val declare_fixed = map_consts o Symtab.delete_safe; val declare_const = map_consts o Symtab.update; (** bounds **) fun next_bound (a, T) ctxt = let val b = Name.bound (#1 (#bounds (rep_data ctxt))); val ctxt' = ctxt |> map_bounds (fn (next, bounds) => (next + 1, ((b, T), a) :: bounds)); in (Free (b, T), ctxt') end; fun revert_bounds ctxt t = (case #2 (#bounds (rep_data ctxt)) of [] => t | bounds => let val names = Term.declare_term_names t (names_of ctxt); val xs = rev (#1 (fold_map Name.variant (rev (map #2 bounds)) names)); fun subst ((b, T), _) x' = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T)); in Term.subst_atomic (map2 subst bounds xs) t end); (** fixes **) (* inner body mode *) val inner_body = Config.declare_bool ("inner_body", \<^here>) (K false); val is_body = Config.apply inner_body; val set_body = Config.put inner_body; val restore_body = set_body o is_body; (* proper mode *) val proper_fixes = Config.declare_bool ("proper_fixes", \<^here>) (K true); val improper_fixes = Config.put proper_fixes false; val restore_proper_fixes = Config.put proper_fixes o Config.apply proper_fixes; fun is_improper ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (_, proper) => not proper | NONE => false); (* specialized name space *) val is_fixed = Name_Space.defined o fixes_of; fun is_newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer); val fixed_ord = Name_Space.entry_ord o fixes_space; val intern_fixed = Name_Space.intern o fixes_space; fun lookup_fixed ctxt x = let val x' = intern_fixed ctxt x in if is_fixed ctxt x' then SOME x' else NONE end; fun revert_fixed ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (x', _) => if intern_fixed ctxt x' = x then x' else x | NONE => x); fun markup_fixed ctxt x = Name_Space.markup (fixes_space ctxt) x |> Markup.name (revert_fixed ctxt x); fun markup ctxt x = if is_improper ctxt x then Markup.improper else if Name.is_skolem x then Markup.skolem else Markup.free; val markup_entity_def = Name_Space.markup_def o fixes_space; fun dest_fixes ctxt = Name_Space.fold_table (fn (x, (y, _)) => cons (y, x)) (fixes_of ctxt) [] |> sort (Name_Space.entry_ord (fixes_space ctxt) o apply2 #2); (* collect variables *) fun add_free_names ctxt = fold_aterms (fn Free (x, _) => not (is_fixed ctxt x) ? insert (op =) x | _ => I); fun add_frees ctxt = fold_aterms (fn Free (x, T) => not (is_fixed ctxt x) ? insert (op =) (x, T) | _ => I); fun add_fixed_names ctxt = fold_aterms (fn Free (x, _) => is_fixed ctxt x ? insert (op =) x | _ => I); fun add_fixed ctxt = fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I); fun add_newly_fixed ctxt' ctxt = fold_aterms (fn Free (x, T) => is_newly_fixed ctxt' ctxt x ? insert (op =) (x, T) | _ => I); (* declarations *) local fun err_dups dups = error ("Duplicate fixed variable(s): " ^ commas (map Binding.print dups)); fun new_fixed ((x, x'), pos) ctxt = if is_some (lookup_fixed ctxt x') then err_dups [Binding.make (x, pos)] else let val proper = Config.get ctxt proper_fixes; val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming) |> Context_Position.set_visible_generic false; in ctxt |> map_fixes (Name_Space.define context true (Binding.make (x', pos), (x, proper)) #> snd #> Name_Space.alias_table Name_Space.global_naming (Binding.make (x, pos)) x') |> declare_fixed x |> declare_constraints (Syntax.free x') end; fun new_fixes names' args = map_names (K names') #> fold new_fixed args #> pair (map (#2 o #1) args); in fun add_fixes_binding bs ctxt = let val _ = (case filter (Name.is_skolem o Binding.name_of) bs of [] => () | bads => error ("Illegal internal Skolem constant(s): " ^ commas (map Binding.print bads))); val _ = (case duplicates (op = o apply2 Binding.name_of) bs of [] => () | dups => err_dups dups); val xs = map check_name bs; val names = names_of ctxt; val (xs', names') = if is_body ctxt then fold_map Name.variant xs names |>> map Name.skolem else (xs, fold Name.declare xs names); in ctxt |> new_fixes names' ((xs ~~ xs') ~~ map Binding.pos_of bs) end; fun variant_names ctxt raw_xs = let val names = names_of ctxt; val xs = map (fn x => Name.clean x |> Name.is_internal x ? Name.internal) raw_xs; val (xs', names') = fold_map Name.variant xs names |>> (is_body ctxt ? map Name.skolem); in (names', xs ~~ xs') end; fun variant_fixes xs ctxt = let val (names', vs) = variant_names ctxt xs; in ctxt |> new_fixes names' (map (rpair Position.none) vs) end; fun bound_fixes xs ctxt = let val (names', vs) = variant_names ctxt (map #1 xs); val (ys, ctxt') = fold_map next_bound (map2 (fn (x', _) => fn (_, T) => (x', T)) vs xs) ctxt; val fixes = map2 (fn (x, _) => fn Free (y, _) => ((x, y), Position.none)) vs ys; in ctxt' |> new_fixes names' fixes end; end; val add_fixes = add_fixes_binding o map Binding.name; fun add_fixes_direct xs ctxt = ctxt |> set_body false |> (snd o add_fixes xs) |> restore_body ctxt; fun add_fixes_implicit t ctxt = ctxt |> not (is_body ctxt) ? add_fixes_direct (rev (add_free_names ctxt t [])); (* dummy patterns *) fun fix_dummy_patterns (Const ("Pure.dummy_pattern", T)) ctxt = let val ([x], ctxt') = ctxt |> set_body true |> add_fixes [Name.uu_] ||> restore_body ctxt in (Free (x, T), ctxt') end | fix_dummy_patterns (Abs (x, T, b)) ctxt = let val (b', ctxt') = fix_dummy_patterns b ctxt in (Abs (x, T, b'), ctxt') end | fix_dummy_patterns (t $ u) ctxt = let val (t', ctxt') = fix_dummy_patterns t ctxt; val (u', ctxt'') = fix_dummy_patterns u ctxt'; in (t' $ u', ctxt'') end | fix_dummy_patterns a ctxt = (a, ctxt); (** export -- generalize type/term variables (beware of closure sizes) **) fun gen_all ctxt th = let val i = Thm.maxidx_thm th (maxidx_of ctxt) + 1; fun gen (x, T) = Thm.forall_elim (Thm.cterm_of ctxt (Var ((x, i), T))); in fold gen (Drule.outer_params (Thm.prop_of th)) th end; fun export_inst inner outer = let val declared_outer = is_declared outer; val still_fixed = not o is_newly_fixed inner outer; val gen_fixes = Name_Space.fold_table (fn (y, _) => not (is_fixed outer y) ? Symtab.insert_set y) (fixes_of inner) Symtab.empty; val type_occs_inner = type_occs_of inner; fun gen_fixesT ts = Symtab.fold (fn (a, xs) => if declared_outer a orelse exists still_fixed xs then I else Symtab.insert_set a) (fold decl_type_occs ts type_occs_inner) Symtab.empty; in (gen_fixesT, gen_fixes) end; fun exportT_inst inner outer = #1 (export_inst inner outer); fun exportT_terms inner outer = let val mk_tfrees = exportT_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, Symtab.empty) (fold (Term.fold_types Term.maxidx_typ) ts maxidx + 1)) end; fun export_terms inner outer = let val (mk_tfrees, tfrees) = export_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, tfrees) (fold Term.maxidx_term ts maxidx + 1)) end; fun export_prf inner outer prf = let val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; val tfrees = mk_tfrees []; val maxidx = maxidx_of outer; val idx = Proofterm.maxidx_proof prf maxidx + 1; val gen_term = Term_Subst.generalize_same (tfrees, frees) idx; val gen_typ = Term_Subst.generalizeT_same tfrees idx; in Same.commit (Proofterm.map_proof_terms_same gen_term gen_typ) prf end; fun gen_export (mk_tfrees, frees) maxidx ths = let val tfrees = mk_tfrees (map Thm.full_prop_of ths); val idx = fold Thm.maxidx_thm ths maxidx + 1; in map (Thm.generalize (tfrees, frees) idx) ths end; fun exportT inner outer = gen_export (exportT_inst inner outer, Symtab.empty) (maxidx_of outer); fun export inner outer = gen_export (export_inst inner outer) (maxidx_of outer); fun export_morphism inner outer = let val fact = export inner outer; val term = singleton (export_terms inner outer); val typ = Logic.type_map term; in Morphism.transfer_morphism' inner $> Morphism.transfer_morphism' outer $> Morphism.morphism "Variable.export" {binding = [], typ = [typ], term = [term], fact = [fact]} end; (** import -- fix schematic type/term variables **) fun invent_types Ss ctxt = let val tfrees = Name.invent (names_of ctxt) Name.aT (length Ss) ~~ Ss; val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; in (tfrees, ctxt') end; fun importT_inst ts ctxt = let val tvars = rev (fold Term.add_tvars ts []); val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; val instT = fold2 (fn a => fn b => Term_Subst.TVars.add (a, TFree b)) tvars tfrees Term_Subst.TVars.empty; in (instT, ctxt') end; fun import_inst is_open ts ctxt = let val ren = Name.clean #> (if is_open then I else Name.internal); val (instT, ctxt') = importT_inst ts ctxt; val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (fold Term.add_vars ts [])); val (ys, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; val inst = fold2 (fn (x, T) => fn y => Term_Subst.Vars.add ((x, T), Free (y, T))) vars ys Term_Subst.Vars.empty; in ((instT, inst), ctxt'') end; fun importT_terms ts ctxt = let val (instT, ctxt') = importT_inst ts ctxt in (map (Term_Subst.instantiate (instT, Term_Subst.Vars.empty)) ts, ctxt') end; fun import_terms is_open ts ctxt = let val (inst, ctxt') = import_inst is_open ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end; fun importT ths ctxt = let val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; val instT' = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) instT []; val ths' = map (Thm.instantiate (instT', [])) ths; in ((instT', ths'), ctxt') end; fun import_prf is_open prf ctxt = let val ts = rev (Proofterm.fold_proof_terms_types cons (cons o Logic.mk_type) prf []); val (insts, ctxt') = import_inst is_open ts ctxt; in (Proofterm.instantiate insts prf, ctxt') end; fun import is_open ths ctxt = let val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; val instT' = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) instT []; val inst' = Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt' t)) inst []; val ths' = map (Thm.instantiate (instT', inst')) ths; in (((instT', inst'), ths'), ctxt') end; fun import_vars ctxt th = let val ((_, [th']), _) = ctxt |> set_body false |> import true [th]; in th' end; (* import/export *) fun gen_trade imp exp f ctxt ths = let val ((_, ths'), ctxt') = imp ths ctxt in exp ctxt' ctxt (f ctxt' ths') end; val tradeT = gen_trade importT exportT; val trade = gen_trade (import true) export; (* focus on outermost parameters: \x y z. B *) val bound_focus = Config.declare_bool ("bound_focus", \<^here>) (K false); val is_bound_focus = Config.apply bound_focus; val set_bound_focus = Config.put bound_focus; val restore_bound_focus = set_bound_focus o is_bound_focus; fun focus_params bindings t ctxt = let val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) val (xs, Ts) = split_list ps; val (xs', ctxt') = (case bindings of SOME bs => ctxt |> set_body true |> add_fixes_binding bs ||> restore_body ctxt | NONE => if is_bound_focus ctxt then bound_fixes ps ctxt else variant_fixes xs ctxt); val ps' = xs' ~~ Ts; val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps'; in ((xs, ps'), ctxt'') end; fun focus bindings t ctxt = let val ((xs, ps), ctxt') = focus_params bindings t ctxt; val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t); in (((xs ~~ ps), t'), ctxt') end; fun forall_elim_prop t prop = Thm.beta_conversion false (Thm.apply (Thm.dest_arg prop) t) |> Thm.cprop_of |> Thm.dest_arg; fun focus_cterm bindings goal ctxt = let val ((xs, ps), ctxt') = focus_params bindings (Thm.term_of goal) ctxt; val ps' = map (Thm.cterm_of ctxt' o Free) ps; val goal' = fold forall_elim_prop ps' goal; in ((xs ~~ ps', goal'), ctxt') end; fun focus_subgoal bindings i st = let - val all_vars = Thm.fold_terms Term.add_vars st []; + val all_vars = Thm.fold_terms Term_Subst.add_vars st Term_Subst.Vars.empty; in - fold (unbind_term o #1) all_vars #> - fold (declare_constraints o Var) all_vars #> + Term_Subst.Vars.fold (unbind_term o #1 o #1) all_vars #> + Term_Subst.Vars.fold (declare_constraints o Var o #1) all_vars #> focus_cterm bindings (Thm.cprem_of st i) end; (** implicit polymorphism **) (* warn_extra_tfrees *) fun warn_extra_tfrees ctxt1 ctxt2 = let fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); fun occs_free a x = (case def_type ctxt1 false (x, ~1) of SOME T => if occs_typ a T then I else cons (a, x) | NONE => cons (a, x)); val occs1 = type_occs_of ctxt1; val occs2 = type_occs_of ctxt2; val extras = Symtab.fold (fn (a, xs) => if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; val tfrees = map #1 extras |> sort_distinct string_ord; val frees = map #2 extras |> sort_distinct string_ord; in if null extras orelse not (Context_Position.is_visible ctxt2) then () else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^ space_implode " or " (map quote frees)) end; (* polymorphic terms *) fun polymorphic_types ctxt ts = let val ctxt' = fold declare_term ts ctxt; val occs = type_occs_of ctxt; val occs' = type_occs_of ctxt'; val types = Symtab.fold (fn (a, _) => if Symtab.defined occs a then I else Symtab.insert_set a) occs' Symtab.empty; val idx = maxidx_of ctxt' + 1; val Ts' = (fold o fold_types o fold_atyps) (fn T as TFree _ => (case Term_Subst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) | _ => I) ts []; val ts' = map (Term_Subst.generalize (types, Symtab.empty) idx) ts; in (rev Ts', ts') end; fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); end;