clasohm@0: (* Title: drule clasohm@0: ID: $Id$ clasohm@0: Author: Lawrence C Paulson, Cambridge University Computer Laboratory clasohm@0: Copyright 1993 University of Cambridge clasohm@0: clasohm@0: Derived rules and other operations on theorems and theories clasohm@0: *) clasohm@0: clasohm@0: infix 0 RS RSN RL RLN COMP; clasohm@0: clasohm@0: signature DRULE = clasohm@0: sig clasohm@0: structure Thm : THM clasohm@0: local open Thm in clasohm@0: val asm_rl: thm clasohm@0: val assume_ax: theory -> string -> thm clasohm@0: val COMP: thm * thm -> thm clasohm@0: val compose: thm * int * thm -> thm list clasohm@0: val cterm_instantiate: (Sign.cterm*Sign.cterm)list -> thm -> thm clasohm@0: val cut_rl: thm clasohm@0: val equal_abs_elim: Sign.cterm -> thm -> thm clasohm@0: val equal_abs_elim_list: Sign.cterm list -> thm -> thm clasohm@0: val eq_sg: Sign.sg * Sign.sg -> bool clasohm@0: val eq_thm: thm * thm -> bool clasohm@0: val eq_thm_sg: thm * thm -> bool clasohm@0: val flexpair_abs_elim_list: Sign.cterm list -> thm -> thm clasohm@0: val forall_intr_list: Sign.cterm list -> thm -> thm clasohm@0: val forall_intr_frees: thm -> thm clasohm@0: val forall_elim_list: Sign.cterm list -> thm -> thm clasohm@0: val forall_elim_var: int -> thm -> thm clasohm@0: val forall_elim_vars: int -> thm -> thm clasohm@0: val implies_elim_list: thm -> thm list -> thm clasohm@0: val implies_intr_list: Sign.cterm list -> thm -> thm clasohm@0: val print_cterm: Sign.cterm -> unit clasohm@0: val print_ctyp: Sign.ctyp -> unit clasohm@0: val print_goals: int -> thm -> unit clasohm@0: val print_sg: Sign.sg -> unit clasohm@0: val print_theory: theory -> unit clasohm@0: val pprint_sg: Sign.sg -> pprint_args -> unit clasohm@0: val pprint_theory: theory -> pprint_args -> unit clasohm@0: val print_thm: thm -> unit clasohm@0: val prth: thm -> thm clasohm@0: val prthq: thm Sequence.seq -> thm Sequence.seq clasohm@0: val prths: thm list -> thm list clasohm@0: val read_instantiate: (string*string)list -> thm -> thm clasohm@0: val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm clasohm@0: val reflexive_thm: thm clasohm@0: val revcut_rl: thm clasohm@0: val rewrite_goal_rule: (meta_simpset -> thm -> thm option) -> meta_simpset -> clasohm@0: int -> thm -> thm clasohm@0: val rewrite_goals_rule: thm list -> thm -> thm clasohm@0: val rewrite_rule: thm list -> thm -> thm clasohm@0: val RS: thm * thm -> thm clasohm@0: val RSN: thm * (int * thm) -> thm clasohm@0: val RL: thm list * thm list -> thm list clasohm@0: val RLN: thm list * (int * thm list) -> thm list clasohm@0: val show_hyps: bool ref clasohm@0: val size_of_thm: thm -> int clasohm@0: val standard: thm -> thm clasohm@0: val string_of_thm: thm -> string clasohm@0: val symmetric_thm: thm clasohm@0: val pprint_thm: thm -> pprint_args -> unit clasohm@0: val transitive_thm: thm clasohm@0: val triv_forall_equality: thm clasohm@0: val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option) clasohm@0: val zero_var_indexes: thm -> thm clasohm@0: end clasohm@0: end; clasohm@0: clasohm@0: functor DruleFun (structure Logic: LOGIC and Thm: THM) : DRULE = clasohm@0: struct clasohm@0: structure Thm = Thm; clasohm@0: structure Sign = Thm.Sign; clasohm@0: structure Type = Sign.Type; clasohm@0: structure Pretty = Sign.Syntax.Pretty clasohm@0: local open Thm clasohm@0: in clasohm@0: clasohm@0: (**** More derived rules and operations on theorems ****) clasohm@0: clasohm@0: (*** Find the type (sort) associated with a (T)Var or (T)Free in a term clasohm@0: Used for establishing default types (of variables) and sorts (of clasohm@0: type variables) when reading another term. clasohm@0: Index -1 indicates that a (T)Free rather than a (T)Var is wanted. clasohm@0: ***) clasohm@0: clasohm@0: fun types_sorts thm = clasohm@0: let val {prop,hyps,...} = rep_thm thm; clasohm@0: val big = list_comb(prop,hyps); (* bogus term! *) clasohm@0: val vars = map dest_Var (term_vars big); clasohm@0: val frees = map dest_Free (term_frees big); clasohm@0: val tvars = term_tvars big; clasohm@0: val tfrees = term_tfrees big; clasohm@0: fun typ(a,i) = if i<0 then assoc(frees,a) else assoc(vars,(a,i)); clasohm@0: fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i)); clasohm@0: in (typ,sort) end; clasohm@0: clasohm@0: (** Standardization of rules **) clasohm@0: clasohm@0: (*Generalization over a list of variables, IGNORING bad ones*) clasohm@0: fun forall_intr_list [] th = th clasohm@0: | forall_intr_list (y::ys) th = clasohm@0: let val gth = forall_intr_list ys th clasohm@0: in forall_intr y gth handle THM _ => gth end; clasohm@0: clasohm@0: (*Generalization over all suitable Free variables*) clasohm@0: fun forall_intr_frees th = clasohm@0: let val {prop,sign,...} = rep_thm th clasohm@0: in forall_intr_list clasohm@0: (map (Sign.cterm_of sign) (sort atless (term_frees prop))) clasohm@0: th clasohm@0: end; clasohm@0: clasohm@0: (*Replace outermost quantified variable by Var of given index. clasohm@0: Could clash with Vars already present.*) clasohm@0: fun forall_elim_var i th = clasohm@0: let val {prop,sign,...} = rep_thm th clasohm@0: in case prop of clasohm@0: Const("all",_) $ Abs(a,T,_) => clasohm@0: forall_elim (Sign.cterm_of sign (Var((a,i), T))) th clasohm@0: | _ => raise THM("forall_elim_var", i, [th]) clasohm@0: end; clasohm@0: clasohm@0: (*Repeat forall_elim_var until all outer quantifiers are removed*) clasohm@0: fun forall_elim_vars i th = clasohm@0: forall_elim_vars i (forall_elim_var i th) clasohm@0: handle THM _ => th; clasohm@0: clasohm@0: (*Specialization over a list of cterms*) clasohm@0: fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th); clasohm@0: clasohm@0: (* maps [A1,...,An], B to [| A1;...;An |] ==> B *) clasohm@0: fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th); clasohm@0: clasohm@0: (* maps [| A1;...;An |] ==> B and [A1,...,An] to B *) clasohm@0: fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths); clasohm@0: clasohm@0: (*Reset Var indexes to zero, renaming to preserve distinctness*) clasohm@0: fun zero_var_indexes th = clasohm@0: let val {prop,sign,...} = rep_thm th; clasohm@0: val vars = term_vars prop clasohm@0: val bs = foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars) clasohm@0: val inrs = add_term_tvars(prop,[]); clasohm@0: val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs)); clasohm@0: val tye = map (fn ((v,rs),a) => (v, TVar((a,0),rs))) (inrs ~~ nms') clasohm@0: val ctye = map (fn (v,T) => (v,Sign.ctyp_of sign T)) tye; clasohm@0: fun varpairs([],[]) = [] clasohm@0: | varpairs((var as Var(v,T)) :: vars, b::bs) = clasohm@0: let val T' = typ_subst_TVars tye T clasohm@0: in (Sign.cterm_of sign (Var(v,T')), clasohm@0: Sign.cterm_of sign (Var((b,0),T'))) :: varpairs(vars,bs) clasohm@0: end clasohm@0: | varpairs _ = raise TERM("varpairs", []); clasohm@0: in instantiate (ctye, varpairs(vars,rev bs)) th end; clasohm@0: clasohm@0: clasohm@0: (*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers; clasohm@0: all generality expressed by Vars having index 0.*) clasohm@0: fun standard th = clasohm@0: let val {maxidx,...} = rep_thm th clasohm@0: in varifyT (zero_var_indexes (forall_elim_vars(maxidx+1) clasohm@0: (forall_intr_frees(implies_intr_hyps th)))) clasohm@0: end; clasohm@0: clasohm@0: (*Assume a new formula, read following the same conventions as axioms. clasohm@0: Generalizes over Free variables, clasohm@0: creates the assumption, and then strips quantifiers. clasohm@0: Example is [| ALL x:?A. ?P(x) |] ==> [| ?P(?a) |] clasohm@0: [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ] *) clasohm@0: fun assume_ax thy sP = clasohm@0: let val sign = sign_of thy clasohm@0: val prop = Logic.close_form (Sign.term_of (Sign.read_cterm sign clasohm@0: (sP, propT))) clasohm@0: in forall_elim_vars 0 (assume (Sign.cterm_of sign prop)) end; clasohm@0: clasohm@0: (*Resolution: exactly one resolvent must be produced.*) clasohm@0: fun tha RSN (i,thb) = clasohm@0: case Sequence.chop (2, biresolution false [(false,tha)] i thb) of clasohm@0: ([th],_) => th clasohm@0: | ([],_) => raise THM("RSN: no unifiers", i, [tha,thb]) clasohm@0: | _ => raise THM("RSN: multiple unifiers", i, [tha,thb]); clasohm@0: clasohm@0: (*resolution: P==>Q, Q==>R gives P==>R. *) clasohm@0: fun tha RS thb = tha RSN (1,thb); clasohm@0: clasohm@0: (*For joining lists of rules*) clasohm@0: fun thas RLN (i,thbs) = clasohm@0: let val resolve = biresolution false (map (pair false) thas) i clasohm@0: fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => [] clasohm@0: in flat (map resb thbs) end; clasohm@0: clasohm@0: fun thas RL thbs = thas RLN (1,thbs); clasohm@0: clasohm@0: (*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R clasohm@0: with no lifting or renaming! Q may contain ==> or meta-quants clasohm@0: ALWAYS deletes premise i *) clasohm@0: fun compose(tha,i,thb) = clasohm@0: Sequence.list_of_s (bicompose false (false,tha,0) i thb); clasohm@0: clasohm@0: (*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*) clasohm@0: fun tha COMP thb = clasohm@0: case compose(tha,1,thb) of clasohm@0: [th] => th clasohm@0: | _ => raise THM("COMP", 1, [tha,thb]); clasohm@0: clasohm@0: (*Instantiate theorem th, reading instantiations under signature sg*) clasohm@0: fun read_instantiate_sg sg sinsts th = clasohm@0: let val ts = types_sorts th; clasohm@0: val instpair = Sign.read_insts sg ts ts sinsts clasohm@0: in instantiate instpair th end; clasohm@0: clasohm@0: (*Instantiate theorem th, reading instantiations under theory of th*) clasohm@0: fun read_instantiate sinsts th = clasohm@0: read_instantiate_sg (#sign (rep_thm th)) sinsts th; clasohm@0: clasohm@0: clasohm@0: (*Left-to-right replacements: tpairs = [...,(vi,ti),...]. clasohm@0: Instantiates distinct Vars by terms, inferring type instantiations. *) clasohm@0: local clasohm@0: fun add_types ((ct,cu), (sign,tye)) = clasohm@0: let val {sign=signt, t=t, T= T, ...} = Sign.rep_cterm ct clasohm@0: and {sign=signu, t=u, T= U, ...} = Sign.rep_cterm cu clasohm@0: val sign' = Sign.merge(sign, Sign.merge(signt, signu)) clasohm@0: val tye' = Type.unify (#tsig(Sign.rep_sg sign')) ((T,U), tye) clasohm@0: handle Type.TUNIFY => raise TYPE("add_types", [T,U], [t,u]) clasohm@0: in (sign', tye') end; clasohm@0: in clasohm@0: fun cterm_instantiate ctpairs0 th = clasohm@0: let val (sign,tye) = foldr add_types (ctpairs0, (#sign(rep_thm th),[])) clasohm@0: val tsig = #tsig(Sign.rep_sg sign); clasohm@0: fun instT(ct,cu) = let val inst = subst_TVars tye clasohm@0: in (Sign.cfun inst ct, Sign.cfun inst cu) end clasohm@0: fun ctyp2 (ix,T) = (ix, Sign.ctyp_of sign T) clasohm@0: in instantiate (map ctyp2 tye, map instT ctpairs0) th end clasohm@0: handle TERM _ => clasohm@0: raise THM("cterm_instantiate: incompatible signatures",0,[th]) clasohm@0: | TYPE _ => raise THM("cterm_instantiate: types", 0, [th]) clasohm@0: end; clasohm@0: clasohm@0: clasohm@0: (*** Printing of theorems ***) clasohm@0: clasohm@0: (*If false, hypotheses are printed as dots*) clasohm@0: val show_hyps = ref true; clasohm@0: clasohm@0: fun pretty_thm th = clasohm@0: let val {sign, hyps, prop,...} = rep_thm th clasohm@0: val hsymbs = if null hyps then [] clasohm@0: else if !show_hyps then clasohm@0: [Pretty.brk 2, clasohm@0: Pretty.lst("[","]") (map (Sign.pretty_term sign) hyps)] clasohm@0: else Pretty.str" [" :: map (fn _ => Pretty.str".") hyps @ clasohm@0: [Pretty.str"]"]; clasohm@0: in Pretty.blk(0, Sign.pretty_term sign prop :: hsymbs) end; clasohm@0: clasohm@0: val string_of_thm = Pretty.string_of o pretty_thm; clasohm@0: clasohm@0: val pprint_thm = Pretty.pprint o Pretty.quote o pretty_thm; clasohm@0: clasohm@0: clasohm@0: (** Top-level commands for printing theorems **) clasohm@0: val print_thm = writeln o string_of_thm; clasohm@0: clasohm@0: fun prth th = (print_thm th; th); clasohm@0: clasohm@0: (*Print and return a sequence of theorems, separated by blank lines. *) clasohm@0: fun prthq thseq = clasohm@0: (Sequence.prints (fn _ => print_thm) 100000 thseq; clasohm@0: thseq); clasohm@0: clasohm@0: (*Print and return a list of theorems, separated by blank lines. *) clasohm@0: fun prths ths = (print_list_ln print_thm ths; ths); clasohm@0: clasohm@0: (*Other printing commands*) clasohm@0: val print_cterm = writeln o Sign.string_of_cterm; clasohm@0: val print_ctyp = writeln o Sign.string_of_ctyp; clasohm@0: fun pretty_sg sg = clasohm@0: Pretty.lst ("{", "}") (map (Pretty.str o !) (#stamps (Sign.rep_sg sg))); clasohm@0: clasohm@0: val pprint_sg = Pretty.pprint o pretty_sg; clasohm@0: clasohm@0: val pprint_theory = pprint_sg o sign_of; clasohm@0: clasohm@0: val print_sg = writeln o Pretty.string_of o pretty_sg; clasohm@0: val print_theory = print_sg o sign_of; clasohm@0: clasohm@0: clasohm@0: (** Print thm A1,...,An/B in "goal style" -- premises as numbered subgoals **) clasohm@0: clasohm@0: fun prettyprints es = writeln(Pretty.string_of(Pretty.blk(0,es))); clasohm@0: clasohm@0: fun print_goals maxgoals th : unit = clasohm@0: let val {sign, hyps, prop,...} = rep_thm th; clasohm@0: fun printgoals (_, []) = () clasohm@0: | printgoals (n, A::As) = clasohm@0: let val prettyn = Pretty.str(" " ^ string_of_int n ^ ". "); clasohm@0: val prettyA = Sign.pretty_term sign A clasohm@0: in prettyprints[prettyn,prettyA]; clasohm@0: printgoals (n+1,As) clasohm@0: end; clasohm@0: fun prettypair(t,u) = clasohm@0: Pretty.blk(0, [Sign.pretty_term sign t, Pretty.str" =?=", Pretty.brk 1, clasohm@0: Sign.pretty_term sign u]); clasohm@0: fun printff [] = () clasohm@0: | printff tpairs = clasohm@0: writeln("\nFlex-flex pairs:\n" ^ clasohm@0: Pretty.string_of(Pretty.lst("","") (map prettypair tpairs))) clasohm@0: val (tpairs,As,B) = Logic.strip_horn(prop); clasohm@0: val ngoals = length As clasohm@0: in clasohm@0: writeln (Sign.string_of_term sign B); clasohm@0: if ngoals=0 then writeln"No subgoals!" clasohm@0: else if ngoals>maxgoals clasohm@0: then (printgoals (1, take(maxgoals,As)); clasohm@0: writeln("A total of " ^ string_of_int ngoals ^ " subgoals...")) clasohm@0: else printgoals (1, As); clasohm@0: printff tpairs clasohm@0: end; clasohm@0: clasohm@0: clasohm@0: (** theorem equality test is exported and used by BEST_FIRST **) clasohm@0: clasohm@0: (*equality of signatures means exact identity -- by ref equality*) clasohm@0: fun eq_sg (sg1,sg2) = (#stamps(Sign.rep_sg sg1) = #stamps(Sign.rep_sg sg2)); clasohm@0: clasohm@0: (*equality of theorems uses equality of signatures and clasohm@0: the a-convertible test for terms*) clasohm@0: fun eq_thm (th1,th2) = clasohm@0: let val {sign=sg1, hyps=hyps1, prop=prop1, ...} = rep_thm th1 clasohm@0: and {sign=sg2, hyps=hyps2, prop=prop2, ...} = rep_thm th2 clasohm@0: in eq_sg (sg1,sg2) andalso clasohm@0: aconvs(hyps1,hyps2) andalso clasohm@0: prop1 aconv prop2 clasohm@0: end; clasohm@0: clasohm@0: (*Do the two theorems have the same signature?*) clasohm@0: fun eq_thm_sg (th1,th2) = eq_sg(#sign(rep_thm th1), #sign(rep_thm th2)); clasohm@0: clasohm@0: (*Useful "distance" function for BEST_FIRST*) clasohm@0: val size_of_thm = size_of_term o #prop o rep_thm; clasohm@0: clasohm@0: clasohm@0: (*** Meta-Rewriting Rules ***) clasohm@0: clasohm@0: clasohm@0: val reflexive_thm = clasohm@0: let val cx = Sign.cterm_of Sign.pure (Var(("x",0),TVar(("'a",0),["logic"]))) clasohm@0: in Thm.reflexive cx end; clasohm@0: clasohm@0: val symmetric_thm = clasohm@0: let val xy = Sign.read_cterm Sign.pure ("x::'a::logic == y",propT) clasohm@0: in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end; clasohm@0: clasohm@0: val transitive_thm = clasohm@0: let val xy = Sign.read_cterm Sign.pure ("x::'a::logic == y",propT) clasohm@0: val yz = Sign.read_cterm Sign.pure ("y::'a::logic == z",propT) clasohm@0: val xythm = Thm.assume xy and yzthm = Thm.assume yz clasohm@0: in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end; clasohm@0: clasohm@0: clasohm@0: (** Below, a "conversion" has type sign->term->thm **) clasohm@0: clasohm@0: (*In [A1,...,An]==>B, rewrite the selected A's only -- for rewrite_goals_tac*) clasohm@0: fun goals_conv pred cv sign = clasohm@0: let val triv = reflexive o Sign.cterm_of sign clasohm@0: fun gconv i t = clasohm@0: let val (A,B) = Logic.dest_implies t clasohm@0: val thA = if (pred i) then (cv sign A) else (triv A) clasohm@0: in combination (combination (triv implies) thA) clasohm@0: (gconv (i+1) B) clasohm@0: end clasohm@0: handle TERM _ => triv t clasohm@0: in gconv 1 end; clasohm@0: clasohm@0: (*Use a conversion to transform a theorem*) clasohm@0: fun fconv_rule cv th = clasohm@0: let val {sign,prop,...} = rep_thm th clasohm@0: in equal_elim (cv sign prop) th end; clasohm@0: clasohm@0: (*rewriting conversion*) clasohm@0: fun rew_conv prover mss sign t = clasohm@0: rewrite_cterm mss prover (Sign.cterm_of sign t); clasohm@0: clasohm@0: (*Rewrite a theorem*) clasohm@0: fun rewrite_rule thms = fconv_rule (rew_conv (K(K None)) (Thm.mss_of thms)); clasohm@0: clasohm@0: (*Rewrite the subgoals of a proof state (represented by a theorem) *) clasohm@0: fun rewrite_goals_rule thms = clasohm@0: fconv_rule (goals_conv (K true) (rew_conv (K(K None)) (Thm.mss_of thms))); clasohm@0: clasohm@0: (*Rewrite the subgoal of a proof state (represented by a theorem) *) clasohm@0: fun rewrite_goal_rule prover mss i = clasohm@0: fconv_rule (goals_conv (fn j => j=i) (rew_conv prover mss)); clasohm@0: clasohm@0: clasohm@0: (** Derived rules mainly for METAHYPS **) clasohm@0: clasohm@0: (*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*) clasohm@0: fun equal_abs_elim ca eqth = clasohm@0: let val {sign=signa, t=a, ...} = Sign.rep_cterm ca clasohm@0: and combth = combination eqth (reflexive ca) clasohm@0: val {sign,prop,...} = rep_thm eqth clasohm@0: val (abst,absu) = Logic.dest_equals prop clasohm@0: val cterm = Sign.cterm_of (Sign.merge (sign,signa)) clasohm@0: in transitive (symmetric (beta_conversion (cterm (abst$a)))) clasohm@0: (transitive combth (beta_conversion (cterm (absu$a)))) clasohm@0: end clasohm@0: handle THM _ => raise THM("equal_abs_elim", 0, [eqth]); clasohm@0: clasohm@0: (*Calling equal_abs_elim with multiple terms*) clasohm@0: fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th); clasohm@0: clasohm@0: local clasohm@0: open Logic clasohm@0: val alpha = TVar(("'a",0), []) (* type ?'a::{} *) clasohm@0: fun err th = raise THM("flexpair_inst: ", 0, [th]) clasohm@0: fun flexpair_inst def th = clasohm@0: let val {prop = Const _ $ t $ u, sign,...} = rep_thm th clasohm@0: val cterm = Sign.cterm_of sign clasohm@0: fun cvar a = cterm(Var((a,0),alpha)) clasohm@0: val def' = cterm_instantiate [(cvar"t", cterm t), (cvar"u", cterm u)] clasohm@0: def clasohm@0: in equal_elim def' th clasohm@0: end clasohm@0: handle THM _ => err th | bind => err th clasohm@0: in clasohm@0: val flexpair_intr = flexpair_inst (symmetric flexpair_def) clasohm@0: and flexpair_elim = flexpair_inst flexpair_def clasohm@0: end; clasohm@0: clasohm@0: (*Version for flexflex pairs -- this supports lifting.*) clasohm@0: fun flexpair_abs_elim_list cts = clasohm@0: flexpair_intr o equal_abs_elim_list cts o flexpair_elim; clasohm@0: clasohm@0: clasohm@0: (*** Some useful meta-theorems ***) clasohm@0: clasohm@0: (*The rule V/V, obtains assumption solving for eresolve_tac*) clasohm@0: val asm_rl = trivial(Sign.read_cterm Sign.pure ("PROP ?psi",propT)); clasohm@0: clasohm@0: (*Meta-level cut rule: [| V==>W; V |] ==> W *) clasohm@0: val cut_rl = trivial(Sign.read_cterm Sign.pure clasohm@0: ("PROP ?psi ==> PROP ?theta", propT)); clasohm@0: clasohm@0: (*Generalized elim rule for one conclusion; cut_rl with reversed premises: clasohm@0: [| PROP V; PROP V ==> PROP W |] ==> PROP W *) clasohm@0: val revcut_rl = clasohm@0: let val V = Sign.read_cterm Sign.pure ("PROP V", propT) clasohm@0: and VW = Sign.read_cterm Sign.pure ("PROP V ==> PROP W", propT); clasohm@0: in standard (implies_intr V clasohm@0: (implies_intr VW clasohm@0: (implies_elim (assume VW) (assume V)))) clasohm@0: end; clasohm@0: clasohm@0: (* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*) clasohm@0: val triv_forall_equality = clasohm@0: let val V = Sign.read_cterm Sign.pure ("PROP V", propT) clasohm@0: and QV = Sign.read_cterm Sign.pure ("!!x::'a. PROP V", propT) clasohm@0: and x = Sign.read_cterm Sign.pure ("x", TFree("'a",["logic"])); clasohm@0: in standard (equal_intr (implies_intr QV (forall_elim x (assume QV))) clasohm@0: (implies_intr V (forall_intr x (assume V)))) clasohm@0: end; clasohm@0: clasohm@0: end clasohm@0: end;