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@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@23169: val aconvc : cterm * cterm -> 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@22695: val plain_prop_of: thm -> term wenzelm@22695: val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a 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@22362: val rule_attribute: (Context.generic -> thm -> thm) -> attribute wenzelm@22362: val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute wenzelm@22362: val theory_attributes: attribute list -> theory * thm -> theory * thm wenzelm@22362: val proof_attributes: attribute list -> Proof.context * thm -> Proof.context * thm wenzelm@22362: val no_attributes: 'a -> 'a * 'b list wenzelm@22362: val simple_fact: 'a -> ('a * 'b list) list wenzelm@22682: val read_def_cterms: wenzelm@22682: theory * (indexname -> typ option) * (indexname -> sort option) -> wenzelm@22682: string list -> bool -> (string * typ)list wenzelm@22682: -> cterm list * (indexname * typ)list wenzelm@22682: val read_cterm: theory -> string * typ -> cterm wenzelm@22362: end; wenzelm@22362: wenzelm@22362: structure Thm: THM = wenzelm@22362: struct wenzelm@22362: wenzelm@22695: (** basic operations **) wenzelm@22362: 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@23169: val op aconvc = op aconv o pairself Thm.term_of; wenzelm@23169: wenzelm@22362: fun eq_thm ths = wenzelm@22362: Context.joinable (pairself Thm.theory_of_thm ths) andalso wenzelm@22362: thm_ord ths = EQUAL; wenzelm@22362: wenzelm@22362: val eq_thms = eq_list eq_thm; wenzelm@22362: wenzelm@22362: val eq_thm_thy = 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@22695: (* misc operations *) wenzelm@22695: 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@22682: wenzelm@22682: (** theorem kinds **) wenzelm@22362: wenzelm@22362: val axiomK = "axiom"; wenzelm@22362: val assumptionK = "assumption"; wenzelm@22362: val definitionK = "definition"; wenzelm@22362: val theoremK = "theorem"; wenzelm@22362: val lemmaK = "lemma"; wenzelm@22362: val corollaryK = "corollary"; wenzelm@22362: val internalK = "internal"; wenzelm@22362: wenzelm@22362: wenzelm@22682: wenzelm@22682: (** attributes **) wenzelm@22362: wenzelm@22362: fun rule_attribute f (x, th) = (x, f x th); wenzelm@22362: fun declaration_attribute f (x, th) = (f th x, th); wenzelm@22362: wenzelm@22362: fun apply_attributes mk dest = wenzelm@22362: let wenzelm@22362: fun app [] = I wenzelm@22362: | app ((f: attribute) :: fs) = fn (x, th) => f (mk x, th) |>> dest |> app fs; wenzelm@22362: in app end; wenzelm@22362: wenzelm@22362: val theory_attributes = apply_attributes Context.Theory Context.the_theory; wenzelm@22362: val proof_attributes = apply_attributes Context.Proof Context.the_proof; wenzelm@22362: wenzelm@22362: fun no_attributes x = (x, []); wenzelm@22362: fun simple_fact x = [(x, [])]; wenzelm@22362: wenzelm@22362: wenzelm@22682: (** read/certify terms (obsolete) **) (*exception ERROR*) wenzelm@22682: wenzelm@22682: fun read_def_cterms (thy, types, sorts) used freeze sTs = wenzelm@22682: let wenzelm@22682: val (ts', tye) = Sign.read_def_terms (thy, types, sorts) used freeze sTs; wenzelm@22682: val cts = map (Thm.cterm_of thy) ts' wenzelm@22682: handle TYPE (msg, _, _) => error msg wenzelm@22682: | TERM (msg, _) => error msg; wenzelm@22682: in (cts, tye) end; wenzelm@22682: wenzelm@22682: fun read_cterm thy sT = wenzelm@22682: let val ([ct], _) = read_def_cterms (thy, K NONE, K NONE) [] true [sT] wenzelm@22682: in ct end; wenzelm@22682: wenzelm@22682: 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);