wenzelm@22362: (* Title: Pure/more_thm.ML wenzelm@22362: ID: $Id$ wenzelm@22362: Author: Makarius wenzelm@22362: wenzelm@22907: Further operations on type ctyp/cterm/thm, outside the inference kernel. wenzelm@22362: *) wenzelm@22362: wenzelm@23169: infix aconvc; wenzelm@23169: wenzelm@22362: signature THM = wenzelm@22362: sig wenzelm@22362: include THM wenzelm@24948: val aconvc: cterm * cterm -> bool wenzelm@23491: val add_cterm_frees: cterm -> cterm list -> cterm list wenzelm@22907: val mk_binop: cterm -> cterm -> cterm -> cterm wenzelm@22907: val dest_binop: cterm -> cterm * cterm wenzelm@22907: val dest_implies: cterm -> cterm * cterm wenzelm@22907: val dest_equals: cterm -> cterm * cterm wenzelm@22907: val dest_equals_lhs: cterm -> cterm wenzelm@22907: val dest_equals_rhs: cterm -> cterm wenzelm@22907: val lhs_of: thm -> cterm wenzelm@22907: val rhs_of: thm -> cterm wenzelm@22362: val thm_ord: thm * thm -> order wenzelm@23599: val is_reflexive: thm -> bool wenzelm@22362: val eq_thm: thm * thm -> bool wenzelm@22362: val eq_thms: thm list * thm list -> bool wenzelm@22362: val eq_thm_thy: thm * thm -> bool wenzelm@22362: val eq_thm_prop: thm * thm -> bool wenzelm@22362: val equiv_thm: thm * thm -> bool wenzelm@28621: val check_shyps: sort list -> thm -> thm wenzelm@24048: val is_dummy: thm -> bool wenzelm@22695: val plain_prop_of: thm -> term wenzelm@22695: val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a wenzelm@24048: val add_thm: thm -> thm list -> thm list wenzelm@24048: val del_thm: thm -> thm list -> thm list wenzelm@24048: val merge_thms: thm list * thm list -> thm list wenzelm@27866: val elim_implies: thm -> thm -> thm wenzelm@27866: val forall_elim_var: int -> thm -> thm wenzelm@27866: val forall_elim_vars: int -> thm -> thm wenzelm@27866: val unvarify: thm -> thm wenzelm@27866: val close_derivation: thm -> thm wenzelm@28116: val add_axiom: bstring * term -> theory -> thm * theory wenzelm@27866: val add_def: bool -> bool -> bstring * term -> theory -> thm * theory wenzelm@27866: val rule_attribute: (Context.generic -> thm -> thm) -> attribute wenzelm@27866: val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute wenzelm@27866: val theory_attributes: attribute list -> theory * thm -> theory * thm wenzelm@27866: val proof_attributes: attribute list -> Proof.context * thm -> Proof.context * thm wenzelm@27866: val no_attributes: 'a -> 'a * 'b list wenzelm@27866: val simple_fact: 'a -> ('a * 'b list) list wenzelm@28017: val tag_rule: Properties.property -> thm -> thm wenzelm@27866: val untag_rule: string -> thm -> thm wenzelm@28017: val tag: Properties.property -> attribute wenzelm@27866: val untag: string -> attribute wenzelm@27866: val position_of: thm -> Position.T wenzelm@27866: val default_position: Position.T -> thm -> thm wenzelm@27866: val default_position_of: thm -> thm -> thm wenzelm@27866: val has_name_hint: thm -> bool wenzelm@27866: val get_name_hint: thm -> string wenzelm@27866: val put_name_hint: string -> thm -> thm wenzelm@27866: val get_group: thm -> string option wenzelm@27866: val put_group: string -> thm -> thm wenzelm@27866: val group: string -> attribute wenzelm@22362: val axiomK: string wenzelm@22362: val assumptionK: string wenzelm@22362: val definitionK: string wenzelm@22362: val theoremK: string wenzelm@22362: val lemmaK: string wenzelm@22362: val corollaryK: string wenzelm@22362: val internalK: string wenzelm@27866: val has_kind: thm -> bool wenzelm@27866: val get_kind: thm -> string wenzelm@27866: val kind_rule: string -> thm -> thm wenzelm@27866: val kind: string -> attribute wenzelm@27866: val kind_internal: attribute wenzelm@28017: val has_internal: Properties.property list -> bool wenzelm@27866: val is_internal: thm -> bool wenzelm@22362: end; wenzelm@22362: wenzelm@22362: structure Thm: THM = wenzelm@22362: struct wenzelm@22362: wenzelm@22695: (** basic operations **) wenzelm@22362: wenzelm@23491: (* collecting cterms *) wenzelm@23491: wenzelm@23491: val op aconvc = op aconv o pairself Thm.term_of; wenzelm@23491: wenzelm@23491: fun add_cterm_frees ct = wenzelm@23491: let wenzelm@23491: val cert = Thm.cterm_of (Thm.theory_of_cterm ct); wenzelm@23491: val t = Thm.term_of ct; wenzelm@23491: in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (cert v) | _ => I) t end; wenzelm@23491: wenzelm@23491: wenzelm@22907: (* cterm constructors and destructors *) wenzelm@22907: wenzelm@22907: fun mk_binop c a b = Thm.capply (Thm.capply c a) b; wenzelm@22907: fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); wenzelm@22907: wenzelm@22907: fun dest_implies ct = wenzelm@22907: (case Thm.term_of ct of wenzelm@22907: Const ("==>", _) $ _ $ _ => dest_binop ct wenzelm@22907: | _ => raise TERM ("dest_implies", [Thm.term_of ct])); wenzelm@22907: wenzelm@22907: fun dest_equals ct = wenzelm@22907: (case Thm.term_of ct of wenzelm@22907: Const ("==", _) $ _ $ _ => dest_binop ct wenzelm@22907: | _ => raise TERM ("dest_equals", [Thm.term_of ct])); wenzelm@22907: wenzelm@22907: fun dest_equals_lhs ct = wenzelm@22907: (case Thm.term_of ct of wenzelm@22907: Const ("==", _) $ _ $ _ => Thm.dest_arg1 ct wenzelm@22907: | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); wenzelm@22907: wenzelm@22907: fun dest_equals_rhs ct = wenzelm@22907: (case Thm.term_of ct of wenzelm@22907: Const ("==", _) $ _ $ _ => Thm.dest_arg ct wenzelm@22907: | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); wenzelm@22907: wenzelm@22907: val lhs_of = dest_equals_lhs o Thm.cprop_of; wenzelm@22907: val rhs_of = dest_equals_rhs o Thm.cprop_of; wenzelm@22907: wenzelm@22907: wenzelm@22907: (* thm order: ignores theory context! *) wenzelm@22682: wenzelm@22362: fun thm_ord (th1, th2) = wenzelm@22362: let wenzelm@22362: val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1; wenzelm@22362: val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2; wenzelm@22362: in wenzelm@22362: (case Term.fast_term_ord (prop1, prop2) of wenzelm@22362: EQUAL => wenzelm@22362: (case list_ord (prod_ord Term.fast_term_ord Term.fast_term_ord) (tpairs1, tpairs2) of wenzelm@22362: EQUAL => wenzelm@22362: (case list_ord Term.fast_term_ord (hyps1, hyps2) of wenzelm@22362: EQUAL => list_ord Term.sort_ord (shyps1, shyps2) wenzelm@22362: | ord => ord) wenzelm@22362: | ord => ord) wenzelm@22362: | ord => ord) wenzelm@22362: end; wenzelm@22362: wenzelm@22682: wenzelm@22682: (* equality *) wenzelm@22682: wenzelm@23599: fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) wenzelm@23599: handle TERM _ => false; wenzelm@23599: wenzelm@22362: fun eq_thm ths = wenzelm@22362: Context.joinable (pairself Thm.theory_of_thm ths) andalso wenzelm@26653: is_equal (thm_ord ths); wenzelm@22362: wenzelm@22362: val eq_thms = eq_list eq_thm; wenzelm@22362: wenzelm@26665: val eq_thm_thy = Theory.eq_thy o pairself Thm.theory_of_thm; wenzelm@22362: val eq_thm_prop = op aconv o pairself Thm.full_prop_of; wenzelm@22362: wenzelm@22682: wenzelm@22682: (* pattern equivalence *) wenzelm@22682: wenzelm@22362: fun equiv_thm ths = wenzelm@22362: Pattern.equiv (Theory.merge (pairself Thm.theory_of_thm ths)) (pairself Thm.full_prop_of ths); wenzelm@22362: wenzelm@22362: wenzelm@28621: (* sort hypotheses *) wenzelm@28621: wenzelm@28621: fun check_shyps sorts raw_th = wenzelm@28621: let wenzelm@28621: val th = Thm.strip_shyps raw_th; wenzelm@28621: val prt_sort = Syntax.pretty_sort_global (Thm.theory_of_thm th); wenzelm@28621: val pending = Sorts.subtract sorts (Thm.extra_shyps th); wenzelm@28621: in wenzelm@28621: if null pending then th wenzelm@28621: else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: wenzelm@28621: Pretty.brk 1 :: Pretty.commas (map prt_sort pending)))) wenzelm@28621: end; wenzelm@28621: wenzelm@28621: wenzelm@22695: (* misc operations *) wenzelm@22695: wenzelm@24048: fun is_dummy thm = wenzelm@24048: (case try Logic.dest_term (Thm.concl_of thm) of wenzelm@24048: NONE => false wenzelm@24048: | SOME t => Term.is_dummy_pattern t); wenzelm@24048: wenzelm@22695: fun plain_prop_of raw_thm = wenzelm@22695: let wenzelm@22695: val thm = Thm.strip_shyps raw_thm; wenzelm@22695: fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); wenzelm@22695: val {hyps, prop, tpairs, ...} = Thm.rep_thm thm; wenzelm@22695: in wenzelm@22695: if not (null hyps) then wenzelm@22695: err "theorem may not contain hypotheses" wenzelm@22695: else if not (null (Thm.extra_shyps thm)) then wenzelm@22695: err "theorem may not contain sort hypotheses" wenzelm@22695: else if not (null tpairs) then wenzelm@22695: err "theorem may not contain flex-flex pairs" wenzelm@22695: else prop wenzelm@22695: end; wenzelm@22695: wenzelm@22695: fun fold_terms f th = wenzelm@22695: let val {tpairs, prop, hyps, ...} = Thm.rep_thm th wenzelm@22695: in fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps end; wenzelm@22695: wenzelm@22695: wenzelm@24048: (* lists of theorems in canonical order *) wenzelm@24048: wenzelm@24048: val add_thm = update eq_thm_prop; wenzelm@24048: val del_thm = remove eq_thm_prop; wenzelm@24048: val merge_thms = merge eq_thm_prop; wenzelm@24048: wenzelm@24048: wenzelm@22682: wenzelm@24980: (** basic derived rules **) wenzelm@24980: wenzelm@24980: (*Elimination of implication wenzelm@24980: A A ==> B wenzelm@24980: ------------ wenzelm@24980: B wenzelm@24980: *) wenzelm@24980: fun elim_implies thA thAB = Thm.implies_elim thAB thA; wenzelm@24980: wenzelm@26653: wenzelm@26653: (* forall_elim_var(s) *) wenzelm@26653: wenzelm@26653: local wenzelm@26653: wenzelm@26653: fun forall_elim_vars_aux strip_vars i th = wenzelm@26653: let wenzelm@26653: val thy = Thm.theory_of_thm th; wenzelm@26653: val {tpairs, prop, ...} = Thm.rep_thm th; wenzelm@26653: val add_used = Term.fold_aterms wenzelm@26653: (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I); wenzelm@26653: val used = fold (fn (t, u) => add_used t o add_used u) tpairs (add_used prop []); wenzelm@26653: val vars = strip_vars prop; wenzelm@26653: val cvars = (Name.variant_list used (map #1 vars), vars) wenzelm@26653: |> ListPair.map (fn (x, (_, T)) => Thm.cterm_of thy (Var ((x, i), T))); wenzelm@26653: in fold Thm.forall_elim cvars th end; wenzelm@26653: wenzelm@26653: in wenzelm@26653: wenzelm@26653: val forall_elim_vars = forall_elim_vars_aux Term.strip_all_vars; wenzelm@26653: wenzelm@26653: fun forall_elim_var i th = forall_elim_vars_aux wenzelm@26653: (fn Const ("all", _) $ Abs (a, T, _) => [(a, T)] wenzelm@26653: | _ => raise THM ("forall_elim_vars", i, [th])) i th; wenzelm@26653: wenzelm@26653: end; wenzelm@26653: wenzelm@26653: wenzelm@26653: (* unvarify: global schematic variables *) wenzelm@26653: wenzelm@24980: fun unvarify th = wenzelm@24980: let wenzelm@24980: val thy = Thm.theory_of_thm th; wenzelm@24980: val cert = Thm.cterm_of thy; wenzelm@24980: val certT = Thm.ctyp_of thy; wenzelm@24980: wenzelm@24980: val prop = Thm.full_prop_of th; wenzelm@24980: val _ = map Logic.unvarify (prop :: Thm.hyps_of th) wenzelm@24980: handle TERM (msg, _) => raise THM (msg, 0, [th]); wenzelm@24980: wenzelm@24980: val instT0 = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S))); wenzelm@24980: val instT = map (fn (v, T) => (certT (TVar v), certT T)) instT0; wenzelm@24980: val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => wenzelm@24980: let val T' = TermSubst.instantiateT instT0 T wenzelm@24980: in (cert (Var ((a, i), T')), cert (Free ((a, T')))) end); wenzelm@24980: in Thm.instantiate (instT, inst) th end; wenzelm@24980: wenzelm@26653: wenzelm@26653: (* close_derivation *) wenzelm@26653: wenzelm@26628: fun close_derivation thm = wenzelm@26628: if Thm.get_name thm = "" then Thm.put_name "" thm wenzelm@26628: else thm; wenzelm@26628: wenzelm@24980: wenzelm@24980: wenzelm@24980: (** specification primitives **) wenzelm@24980: wenzelm@28116: fun add_axiom (name, prop) thy = wenzelm@24980: let wenzelm@24980: val name' = if name = "" then "axiom_" ^ serial_string () else name; wenzelm@28116: val thy' = thy |> Theory.add_axioms_i [(name', prop)]; haftmann@28965: val axm = unvarify (Thm.axiom thy' (Sign.full_bname thy' name')); wenzelm@28116: in (axm, thy') end; wenzelm@24980: haftmann@25518: fun add_def unchecked overloaded (name, prop) thy = wenzelm@24980: let wenzelm@24980: val tfrees = rev (map TFree (Term.add_tfrees prop [])); wenzelm@24980: val tfrees' = map (fn a => TFree (a, [])) (Name.invents Name.context Name.aT (length tfrees)); wenzelm@24980: val strip_sorts = tfrees ~~ tfrees'; wenzelm@24980: val recover_sorts = map (pairself (Thm.ctyp_of thy o Logic.varifyT)) (tfrees' ~~ tfrees); wenzelm@24980: wenzelm@24980: val prop' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip_sorts))) prop; haftmann@25518: val thy' = Theory.add_defs_i unchecked overloaded [(name, prop')] thy; haftmann@28965: val axm' = Thm.axiom thy' (Sign.full_bname thy' name); wenzelm@24980: val thm = unvarify (Thm.instantiate (recover_sorts, []) axm'); wenzelm@24980: in (thm, thy') end; wenzelm@24980: wenzelm@27866: wenzelm@27866: wenzelm@27866: (** attributes **) wenzelm@27866: wenzelm@27866: fun rule_attribute f (x, th) = (x, f x th); wenzelm@27866: fun declaration_attribute f (x, th) = (f th x, th); wenzelm@27866: wenzelm@27866: fun apply_attributes mk dest = wenzelm@27866: let wenzelm@27866: fun app [] = I wenzelm@27866: | app ((f: attribute) :: fs) = fn (x, th) => f (mk x, th) |>> dest |> app fs; wenzelm@27866: in app end; wenzelm@27866: wenzelm@27866: val theory_attributes = apply_attributes Context.Theory Context.the_theory; wenzelm@27866: val proof_attributes = apply_attributes Context.Proof Context.the_proof; wenzelm@27866: wenzelm@27866: fun no_attributes x = (x, []); wenzelm@27866: fun simple_fact x = [(x, [])]; wenzelm@27866: wenzelm@27866: wenzelm@27866: wenzelm@27866: (*** theorem tags ***) wenzelm@27866: wenzelm@27866: (* add / delete tags *) wenzelm@27866: wenzelm@27866: fun tag_rule tg = Thm.map_tags (insert (op =) tg); wenzelm@27866: fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); wenzelm@27866: wenzelm@27866: fun tag tg x = rule_attribute (K (tag_rule tg)) x; wenzelm@27866: fun untag s x = rule_attribute (K (untag_rule s)) x; wenzelm@27866: wenzelm@27866: wenzelm@27866: (* position *) wenzelm@27866: wenzelm@27866: val position_of = Position.of_properties o Thm.get_tags; wenzelm@27866: wenzelm@27866: fun default_position pos = Thm.map_tags (Position.default_properties pos); wenzelm@27866: val default_position_of = default_position o position_of; wenzelm@27866: wenzelm@27866: wenzelm@27866: (* unofficial theorem names *) wenzelm@27866: wenzelm@27866: fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); wenzelm@27866: wenzelm@27866: val has_name_hint = can the_name_hint; wenzelm@27866: val get_name_hint = the_default "??.unknown" o try the_name_hint; wenzelm@27866: wenzelm@27866: fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); wenzelm@27866: wenzelm@27866: wenzelm@27866: (* theorem groups *) wenzelm@27866: wenzelm@28017: fun get_group thm = Properties.get (Thm.get_tags thm) Markup.groupN; wenzelm@27866: wenzelm@28017: fun put_group name = if name = "" then I else Thm.map_tags (Properties.put (Markup.groupN, name)); wenzelm@27866: wenzelm@27866: fun group name = rule_attribute (K (put_group name)); wenzelm@27866: wenzelm@27866: wenzelm@27866: (* theorem kinds *) wenzelm@27866: wenzelm@27866: val axiomK = "axiom"; wenzelm@27866: val assumptionK = "assumption"; wenzelm@27866: val definitionK = "definition"; wenzelm@27866: val theoremK = "theorem"; wenzelm@27866: val lemmaK = "lemma"; wenzelm@27866: val corollaryK = "corollary"; wenzelm@27866: val internalK = Markup.internalK; wenzelm@27866: wenzelm@28017: fun the_kind thm = the (Properties.get (Thm.get_tags thm) Markup.kindN); wenzelm@27866: wenzelm@27866: val has_kind = can the_kind; wenzelm@27866: val get_kind = the_default "" o try the_kind; wenzelm@27866: wenzelm@27866: fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; wenzelm@27866: fun kind k x = if k = "" then x else rule_attribute (K (kind_rule k)) x; wenzelm@27866: fun kind_internal x = kind internalK x; wenzelm@27866: fun has_internal tags = exists (fn tg => tg = (Markup.kindN, internalK)) tags; wenzelm@27866: val is_internal = has_internal o Thm.get_tags; wenzelm@27866: wenzelm@27866: wenzelm@22362: open Thm; wenzelm@22362: wenzelm@22362: end; wenzelm@22362: wenzelm@23170: val op aconvc = Thm.aconvc; wenzelm@23170: wenzelm@22362: structure Thmtab = TableFun(type key = thm val ord = Thm.thm_ord);