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