| author | lcp | 
| Mon, 27 Feb 1995 18:07:59 +0100 | |
| changeset 914 | cae574c09137 | 
| parent 843 | c1a4a4206102 | 
| child 922 | 196ca0973a6d | 
| permissions | -rw-r--r-- | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 1 | (* Title: Pure/drule.ML | 
| 0 | 2 | ID: $Id$ | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Derived rules and other operations on theorems and theories | |
| 7 | *) | |
| 8 | ||
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 9 | infix 0 RS RSN RL RLN MRS MRL COMP; | 
| 0 | 10 | |
| 11 | signature DRULE = | |
| 12 | sig | |
| 13 | structure Thm : THM | |
| 14 | local open Thm in | |
| 668 | 15 | val add_defs : (string * string) list -> theory -> theory | 
| 16 | val add_defs_i : (string * term) list -> theory -> theory | |
| 17 | val asm_rl : thm | |
| 18 | val assume_ax : theory -> string -> thm | |
| 19 | val COMP : thm * thm -> thm | |
| 20 | val compose : thm * int * thm -> thm list | |
| 708 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 21 | val cprems_of : thm -> cterm list | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 22 | val cskip_flexpairs : cterm -> cterm | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 23 | val cstrip_imp_prems : cterm -> cterm list | 
| 668 | 24 | val cterm_instantiate : (cterm*cterm)list -> thm -> thm | 
| 25 | val cut_rl : thm | |
| 26 | val equal_abs_elim : cterm -> thm -> thm | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 27 | val equal_abs_elim_list: cterm list -> thm -> thm | 
| 668 | 28 | val eq_thm : thm * thm -> bool | 
| 29 | val eq_thm_sg : thm * thm -> bool | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 30 | val flexpair_abs_elim_list: cterm list -> thm -> thm | 
| 668 | 31 | val forall_intr_list : cterm list -> thm -> thm | 
| 32 | val forall_intr_frees : thm -> thm | |
| 33 | val forall_elim_list : cterm list -> thm -> thm | |
| 34 | val forall_elim_var : int -> thm -> thm | |
| 35 | val forall_elim_vars : int -> thm -> thm | |
| 36 | val implies_elim_list : thm -> thm list -> thm | |
| 37 | val implies_intr_list : cterm list -> thm -> thm | |
| 38 | val MRL : thm list list * thm list -> thm list | |
| 39 | val MRS : thm list * thm -> thm | |
| 40 | val pprint_cterm : cterm -> pprint_args -> unit | |
| 41 | val pprint_ctyp : ctyp -> pprint_args -> unit | |
| 42 | val pprint_theory : theory -> pprint_args -> unit | |
| 43 | val pprint_thm : thm -> pprint_args -> unit | |
| 44 | val pretty_thm : thm -> Sign.Syntax.Pretty.T | |
| 45 | val print_cterm : cterm -> unit | |
| 46 | val print_ctyp : ctyp -> unit | |
| 47 | val print_goals : int -> thm -> unit | |
| 48 | val print_goals_ref : (int -> thm -> unit) ref | |
| 49 | val print_syntax : theory -> unit | |
| 50 | val print_theory : theory -> unit | |
| 51 | val print_thm : thm -> unit | |
| 52 | val prth : thm -> thm | |
| 53 | val prthq : thm Sequence.seq -> thm Sequence.seq | |
| 54 | val prths : thm list -> thm list | |
| 55 | val read_instantiate : (string*string)list -> thm -> thm | |
| 0 | 56 | val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm | 
| 668 | 57 | val read_insts : | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 58 | Sign.sg -> (indexname -> typ option) * (indexname -> sort option) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 59 | -> (indexname -> typ option) * (indexname -> sort option) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 60 | -> (string*string)list | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 61 | -> (indexname*ctyp)list * (cterm*cterm)list | 
| 668 | 62 | val reflexive_thm : thm | 
| 63 | val revcut_rl : thm | |
| 64 | val rewrite_goal_rule : bool*bool -> (meta_simpset -> thm -> thm option) | |
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 65 | -> meta_simpset -> int -> thm -> thm | 
| 0 | 66 | val rewrite_goals_rule: thm list -> thm -> thm | 
| 668 | 67 | val rewrite_rule : thm list -> thm -> thm | 
| 68 | val RS : thm * thm -> thm | |
| 69 | val RSN : thm * (int * thm) -> thm | |
| 70 | val RL : thm list * thm list -> thm list | |
| 71 | val RLN : thm list * (int * thm list) -> thm list | |
| 72 | val show_hyps : bool ref | |
| 73 | val size_of_thm : thm -> int | |
| 74 | val standard : thm -> thm | |
| 75 | val string_of_cterm : cterm -> string | |
| 76 | val string_of_ctyp : ctyp -> string | |
| 77 | val string_of_thm : thm -> string | |
| 78 | val symmetric_thm : thm | |
| 79 | val thin_rl : thm | |
| 80 | val transitive_thm : thm | |
| 0 | 81 | val triv_forall_equality: thm | 
| 82 | val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option) | |
| 668 | 83 | val zero_var_indexes : thm -> thm | 
| 0 | 84 | end | 
| 85 | end; | |
| 86 | ||
| 668 | 87 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 88 | functor DruleFun (structure Logic: LOGIC and Thm: THM): DRULE = | 
| 0 | 89 | struct | 
| 90 | structure Thm = Thm; | |
| 91 | structure Sign = Thm.Sign; | |
| 92 | structure Type = Sign.Type; | |
| 575 | 93 | structure Syntax = Sign.Syntax; | 
| 94 | structure Pretty = Syntax.Pretty | |
| 400 | 95 | structure Symtab = Sign.Symtab; | 
| 96 | ||
| 0 | 97 | local open Thm | 
| 98 | in | |
| 99 | ||
| 561 | 100 | (**** Extend Theories ****) | 
| 101 | ||
| 102 | (** add constant definitions **) | |
| 103 | ||
| 104 | (* all_axioms_of *) | |
| 105 | ||
| 106 | (*results may contain duplicates!*) | |
| 107 | ||
| 108 | fun ancestry_of thy = | |
| 109 | thy :: flat (map ancestry_of (parents_of thy)); | |
| 110 | ||
| 776 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
708diff
changeset | 111 | val all_axioms_of = | 
| 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
708diff
changeset | 112 | flat o map (Symtab.dest o #new_axioms o rep_theory) o ancestry_of; | 
| 561 | 113 | |
| 114 | ||
| 115 | (* clash_types, clash_consts *) | |
| 116 | ||
| 117 | (*check if types have common instance (ignoring sorts)*) | |
| 118 | ||
| 119 | fun clash_types ty1 ty2 = | |
| 120 | let | |
| 121 | val ty1' = Type.varifyT ty1; | |
| 122 | val ty2' = incr_tvar (maxidx_of_typ ty1' + 1) (Type.varifyT ty2); | |
| 123 | in | |
| 124 | Type.raw_unify (ty1', ty2') | |
| 125 | end; | |
| 126 | ||
| 127 | fun clash_consts (c1, ty1) (c2, ty2) = | |
| 128 | c1 = c2 andalso clash_types ty1 ty2; | |
| 129 | ||
| 130 | ||
| 131 | (* clash_defns *) | |
| 132 | ||
| 133 | fun clash_defn c_ty (name, tm) = | |
| 134 | let val (c, ty') = dest_Const (head_of (fst (Logic.dest_equals tm))) in | |
| 135 | if clash_consts c_ty (c, ty') then Some (name, ty') else None | |
| 136 | end handle TERM _ => None; | |
| 137 | ||
| 138 | fun clash_defns c_ty axms = | |
| 139 | distinct (mapfilter (clash_defn c_ty) axms); | |
| 140 | ||
| 141 | ||
| 142 | (* dest_defn *) | |
| 143 | ||
| 144 | fun dest_defn tm = | |
| 145 | let | |
| 146 | fun err msg = raise_term msg [tm]; | |
| 147 | ||
| 148 | val (lhs, rhs) = Logic.dest_equals tm | |
| 149 | handle TERM _ => err "Not a meta-equality (==)"; | |
| 150 | val (head, args) = strip_comb lhs; | |
| 151 | val (c, ty) = dest_Const head | |
| 152 | handle TERM _ => err "Head of lhs not a constant"; | |
| 153 | ||
| 655 | 154 | fun occs_const (Const c_ty') = (c_ty' = (c, ty)) | 
| 561 | 155 | | occs_const (Abs (_, _, t)) = occs_const t | 
| 156 | | occs_const (t $ u) = occs_const t orelse occs_const u | |
| 157 | | occs_const _ = false; | |
| 641 | 158 | |
| 159 | val show_frees = commas_quote o map (fst o dest_Free); | |
| 160 | val show_tfrees = commas_quote o map fst; | |
| 161 | ||
| 162 | val lhs_dups = duplicates args; | |
| 163 | val rhs_extras = gen_rems (op =) (term_frees rhs, args); | |
| 164 | val rhs_extrasT = gen_rems (op =) (term_tfrees rhs, typ_tfrees ty); | |
| 561 | 165 | in | 
| 166 | if not (forall is_Free args) then | |
| 167 | err "Arguments of lhs have to be variables" | |
| 641 | 168 | else if not (null lhs_dups) then | 
| 169 |       err ("Duplicate variables on lhs: " ^ show_frees lhs_dups)
 | |
| 170 | else if not (null rhs_extras) then | |
| 171 |       err ("Extra variables on rhs: " ^ show_frees rhs_extras)
 | |
| 172 | else if not (null rhs_extrasT) then | |
| 173 |       err ("Extra type variables on rhs: " ^ show_tfrees rhs_extrasT)
 | |
| 561 | 174 | else if occs_const rhs then | 
| 655 | 175 |       err ("Constant to be defined occurs on rhs")
 | 
| 561 | 176 | else (c, ty) | 
| 177 | end; | |
| 178 | ||
| 179 | ||
| 180 | (* check_defn *) | |
| 181 | ||
| 641 | 182 | fun err_in_defn name msg = | 
| 183 |   (writeln msg; error ("The error(s) above occurred in definition " ^ quote name));
 | |
| 561 | 184 | |
| 185 | fun check_defn sign (axms, (name, tm)) = | |
| 186 | let | |
| 187 | fun show_const (c, ty) = quote (Pretty.string_of (Pretty.block | |
| 188 | [Pretty.str (c ^ " ::"), Pretty.brk 1, Sign.pretty_typ sign ty])); | |
| 189 | ||
| 190 | fun show_defn c (dfn, ty') = show_const (c, ty') ^ " in " ^ dfn; | |
| 191 | fun show_defns c = commas o map (show_defn c); | |
| 192 | ||
| 193 | val (c, ty) = dest_defn tm | |
| 641 | 194 | handle TERM (msg, _) => err_in_defn name msg; | 
| 561 | 195 | val defns = clash_defns (c, ty) axms; | 
| 196 | in | |
| 197 | if not (null defns) then | |
| 641 | 198 |       err_in_defn name ("Definition of " ^ show_const (c, ty) ^
 | 
| 561 | 199 | " clashes with " ^ show_defns c defns) | 
| 200 | else (name, tm) :: axms | |
| 201 | end; | |
| 202 | ||
| 203 | ||
| 204 | (* add_defs *) | |
| 205 | ||
| 206 | fun ext_defns prep_axm raw_axms thy = | |
| 207 | let | |
| 208 | val axms = map (prep_axm (sign_of thy)) raw_axms; | |
| 209 | val all_axms = all_axioms_of thy; | |
| 210 | in | |
| 211 | foldl (check_defn (sign_of thy)) (all_axms, axms); | |
| 212 | add_axioms_i axms thy | |
| 213 | end; | |
| 214 | ||
| 215 | val add_defs_i = ext_defns cert_axm; | |
| 216 | val add_defs = ext_defns read_axm; | |
| 217 | ||
| 218 | ||
| 219 | ||
| 0 | 220 | (**** More derived rules and operations on theorems ****) | 
| 221 | ||
| 708 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 222 | (** some cterm->cterm operations: much faster than calling cterm_of! **) | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 223 | |
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 224 | (*Discard flexflex pairs; return a cterm*) | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 225 | fun cskip_flexpairs ct = | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 226 | case term_of ct of | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 227 | 	(Const("==>", _) $ (Const("=?=",_)$_$_) $ _) =>
 | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 228 | cskip_flexpairs (#2 (dest_cimplies ct)) | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 229 | | _ => ct; | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 230 | |
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 231 | (* A1==>...An==>B goes to [A1,...,An], where B is not an implication *) | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 232 | fun cstrip_imp_prems ct = | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 233 | let val (cA,cB) = dest_cimplies ct | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 234 | in cA :: cstrip_imp_prems cB end | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 235 | handle TERM _ => []; | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 236 | |
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 237 | (*The premises of a theorem, as a cterm list*) | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 238 | val cprems_of = cstrip_imp_prems o cskip_flexpairs o cprop_of; | 
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 239 | |
| 
8422e50adce0
Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
 lcp parents: 
668diff
changeset | 240 | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 241 | (** reading of instantiations **) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 242 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 243 | fun indexname cs = case Syntax.scan_varname cs of (v,[]) => v | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 244 |         | _ => error("Lexical error in variable name " ^ quote (implode cs));
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 245 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 246 | fun absent ixn = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 247 |   error("No such variable in term: " ^ Syntax.string_of_vname ixn);
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 248 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 249 | fun inst_failure ixn = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 250 |   error("Instantiation of " ^ Syntax.string_of_vname ixn ^ " fails");
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 251 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 252 | fun read_insts sign (rtypes,rsorts) (types,sorts) insts = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 253 | let val {tsig,...} = Sign.rep_sg sign
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 254 | fun split([],tvs,vs) = (tvs,vs) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 255 | | split((sv,st)::l,tvs,vs) = (case explode sv of | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 256 | "'"::cs => split(l,(indexname cs,st)::tvs,vs) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 257 | | cs => split(l,tvs,(indexname cs,st)::vs)); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 258 | val (tvs,vs) = split(insts,[],[]); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 259 | fun readT((a,i),st) = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 260 |         let val ixn = ("'" ^ a,i);
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 261 | val S = case rsorts ixn of Some S => S | None => absent ixn; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 262 | val T = Sign.read_typ (sign,sorts) st; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 263 | in if Type.typ_instance(tsig,T,TVar(ixn,S)) then (ixn,T) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 264 | else inst_failure ixn | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 265 | end | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 266 | val tye = map readT tvs; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 267 | fun add_cterm ((cts,tye), (ixn,st)) = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 268 | let val T = case rtypes ixn of | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 269 | Some T => typ_subst_TVars tye T | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 270 | | None => absent ixn; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 271 | val (ct,tye2) = read_def_cterm (sign,types,sorts) (st,T); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 272 | val cv = cterm_of sign (Var(ixn,typ_subst_TVars tye2 T)) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 273 | in ((cv,ct)::cts,tye2 @ tye) end | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 274 | val (cterms,tye') = foldl add_cterm (([],tye), vs); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 275 | in (map (fn (ixn,T) => (ixn,ctyp_of sign T)) tye', cterms) end; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 276 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 277 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 278 | |
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 279 | (*** Printing of theories, theorems, etc. ***) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 280 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 281 | (*If false, hypotheses are printed as dots*) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 282 | val show_hyps = ref true; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 283 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 284 | fun pretty_thm th = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 285 | let val {sign, hyps, prop,...} = rep_thm th
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 286 | val hsymbs = if null hyps then [] | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 287 | else if !show_hyps then | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 288 | [Pretty.brk 2, | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 289 |                        Pretty.lst("[","]") (map (Sign.pretty_term sign) hyps)]
 | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 290 | else Pretty.str" [" :: map (fn _ => Pretty.str".") hyps @ | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 291 | [Pretty.str"]"]; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 292 | in Pretty.blk(0, Sign.pretty_term sign prop :: hsymbs) end; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 293 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 294 | val string_of_thm = Pretty.string_of o pretty_thm; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 295 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 296 | val pprint_thm = Pretty.pprint o Pretty.quote o pretty_thm; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 297 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 298 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 299 | (** Top-level commands for printing theorems **) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 300 | val print_thm = writeln o string_of_thm; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 301 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 302 | fun prth th = (print_thm th; th); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 303 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 304 | (*Print and return a sequence of theorems, separated by blank lines. *) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 305 | fun prthq thseq = | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 306 | (Sequence.prints (fn _ => print_thm) 100000 thseq; thseq); | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 307 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 308 | (*Print and return a list of theorems, separated by blank lines. *) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 309 | fun prths ths = (print_list_ln print_thm ths; ths); | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 310 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 311 | |
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 312 | (* other printing commands *) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 313 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 314 | fun pprint_ctyp cT = | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 315 |   let val {sign, T} = rep_ctyp cT in Sign.pprint_typ sign T end;
 | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 316 | |
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 317 | fun string_of_ctyp cT = | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 318 |   let val {sign, T} = rep_ctyp cT in Sign.string_of_typ sign T end;
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 319 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 320 | val print_ctyp = writeln o string_of_ctyp; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 321 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 322 | fun pprint_cterm ct = | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 323 |   let val {sign, t, ...} = rep_cterm ct in Sign.pprint_term sign t end;
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 324 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 325 | fun string_of_cterm ct = | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 326 |   let val {sign, t, ...} = rep_cterm ct in Sign.string_of_term sign t end;
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 327 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 328 | val print_cterm = writeln o string_of_cterm; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 329 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 330 | |
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 331 | (* print theory *) | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 332 | |
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 333 | val pprint_theory = Sign.pprint_sg o sign_of; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 334 | |
| 575 | 335 | val print_syntax = Syntax.print_syntax o syn_of; | 
| 336 | ||
| 385 | 337 | fun print_axioms thy = | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 338 | let | 
| 400 | 339 |     val {sign, new_axioms, ...} = rep_theory thy;
 | 
| 340 | val axioms = Symtab.dest new_axioms; | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 341 | |
| 385 | 342 | fun prt_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, | 
| 343 | Pretty.quote (Sign.pretty_term sign t)]; | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 344 | in | 
| 385 | 345 | Pretty.writeln (Pretty.big_list "additional axioms:" (map prt_axm axioms)) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 346 | end; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 347 | |
| 843 | 348 | fun print_theory thy = (Sign.print_sg (sign_of thy); print_axioms thy); | 
| 385 | 349 | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 350 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 351 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 352 | (** Print thm A1,...,An/B in "goal style" -- premises as numbered subgoals **) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 353 | |
| 641 | 354 | (* get type_env, sort_env of term *) | 
| 355 | ||
| 356 | local | |
| 357 | open Syntax; | |
| 358 | ||
| 359 | fun ins_entry (x, y) [] = [(x, [y])] | |
| 360 | | ins_entry (x, y) ((pair as (x', ys')) :: pairs) = | |
| 361 | if x = x' then (x', y ins ys') :: pairs | |
| 362 | else pair :: ins_entry (x, y) pairs; | |
| 363 | ||
| 364 | fun add_type_env (Free (x, T), env) = ins_entry (T, x) env | |
| 365 | | add_type_env (Var (xi, T), env) = ins_entry (T, string_of_vname xi) env | |
| 366 | | add_type_env (Abs (_, _, t), env) = add_type_env (t, env) | |
| 367 | | add_type_env (t $ u, env) = add_type_env (u, add_type_env (t, env)) | |
| 368 | | add_type_env (_, env) = env; | |
| 369 | ||
| 370 | fun add_sort_env (Type (_, Ts), env) = foldr add_sort_env (Ts, env) | |
| 371 | | add_sort_env (TFree (x, S), env) = ins_entry (S, x) env | |
| 372 | | add_sort_env (TVar (xi, S), env) = ins_entry (S, string_of_vname xi) env; | |
| 373 | ||
| 374 | val sort = map (apsnd sort_strings); | |
| 375 | in | |
| 376 | fun type_env t = sort (add_type_env (t, [])); | |
| 377 | fun sort_env t = rev (sort (it_term_types add_sort_env (t, []))); | |
| 378 | end; | |
| 379 | ||
| 380 | ||
| 381 | (* print_goals *) | |
| 382 | ||
| 383 | fun print_goals maxgoals state = | |
| 384 | let | |
| 385 | open Syntax; | |
| 386 | ||
| 387 |     val {sign, prop, ...} = rep_thm state;
 | |
| 388 | ||
| 389 | val pretty_term = Sign.pretty_term sign; | |
| 390 | val pretty_typ = Sign.pretty_typ sign; | |
| 391 | val pretty_sort = Sign.pretty_sort; | |
| 392 | ||
| 393 | fun pretty_vars prtf (X, vs) = Pretty.block | |
| 394 | [Pretty.block (Pretty.commas (map Pretty.str vs)), | |
| 395 | Pretty.str " ::", Pretty.brk 1, prtf X]; | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 396 | |
| 641 | 397 | fun print_list _ _ [] = () | 
| 398 | | print_list name prtf lst = | |
| 399 | (writeln ""; Pretty.writeln (Pretty.big_list name (map prtf lst))); | |
| 400 | ||
| 401 | ||
| 402 | fun print_goals (_, []) = () | |
| 403 | | print_goals (n, A :: As) = (Pretty.writeln (Pretty.blk (0, | |
| 404 |           [Pretty.str (" " ^ string_of_int n ^ ". "), pretty_term A]));
 | |
| 405 | print_goals (n + 1, As)); | |
| 406 | ||
| 407 | val print_ffpairs = | |
| 408 | print_list "Flex-flex pairs:" (pretty_term o Logic.mk_flexpair); | |
| 409 | ||
| 410 | val print_types = print_list "Types:" (pretty_vars pretty_typ) o type_env; | |
| 411 | val print_sorts = print_list "Sorts:" (pretty_vars pretty_sort) o sort_env; | |
| 412 | ||
| 413 | ||
| 414 | val (tpairs, As, B) = Logic.strip_horn prop; | |
| 415 | val ngoals = length As; | |
| 416 | ||
| 417 | val orig_no_freeTs = ! show_no_free_types; | |
| 418 | val orig_sorts = ! show_sorts; | |
| 419 | ||
| 420 | fun restore () = | |
| 421 | (show_no_free_types := orig_no_freeTs; show_sorts := orig_sorts); | |
| 422 | in | |
| 423 | (show_no_free_types := true; show_sorts := false; | |
| 424 | ||
| 425 | Pretty.writeln (pretty_term B); | |
| 426 | ||
| 427 | if ngoals = 0 then writeln "No subgoals!" | |
| 428 | else if ngoals > maxgoals then | |
| 429 | (print_goals (1, take (maxgoals, As)); | |
| 430 |           writeln ("A total of " ^ string_of_int ngoals ^ " subgoals..."))
 | |
| 431 | else print_goals (1, As); | |
| 432 | ||
| 433 | print_ffpairs tpairs; | |
| 434 | ||
| 435 | if orig_sorts then | |
| 436 | (print_types prop; print_sorts prop) | |
| 437 | else if ! show_types then | |
| 438 | print_types prop | |
| 439 | else ()) | |
| 440 | handle exn => (restore (); raise exn); | |
| 441 | restore () | |
| 442 | end; | |
| 443 | ||
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 444 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 445 | (*"hook" for user interfaces: allows print_goals to be replaced*) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 446 | val print_goals_ref = ref print_goals; | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 447 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 448 | (*** Find the type (sort) associated with a (T)Var or (T)Free in a term | 
| 0 | 449 | Used for establishing default types (of variables) and sorts (of | 
| 450 | type variables) when reading another term. | |
| 451 | Index -1 indicates that a (T)Free rather than a (T)Var is wanted. | |
| 452 | ***) | |
| 453 | ||
| 454 | fun types_sorts thm = | |
| 455 |     let val {prop,hyps,...} = rep_thm thm;
 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 456 | val big = list_comb(prop,hyps); (* bogus term! *) | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 457 | val vars = map dest_Var (term_vars big); | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 458 | val frees = map dest_Free (term_frees big); | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 459 | val tvars = term_tvars big; | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 460 | val tfrees = term_tfrees big; | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 461 | fun typ(a,i) = if i<0 then assoc(frees,a) else assoc(vars,(a,i)); | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 462 | fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i)); | 
| 0 | 463 | in (typ,sort) end; | 
| 464 | ||
| 465 | (** Standardization of rules **) | |
| 466 | ||
| 467 | (*Generalization over a list of variables, IGNORING bad ones*) | |
| 468 | fun forall_intr_list [] th = th | |
| 469 | | forall_intr_list (y::ys) th = | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 470 | let val gth = forall_intr_list ys th | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 471 | in forall_intr y gth handle THM _ => gth end; | 
| 0 | 472 | |
| 473 | (*Generalization over all suitable Free variables*) | |
| 474 | fun forall_intr_frees th = | |
| 475 |     let val {prop,sign,...} = rep_thm th
 | |
| 476 | in forall_intr_list | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 477 | (map (cterm_of sign) (sort atless (term_frees prop))) | 
| 0 | 478 | th | 
| 479 | end; | |
| 480 | ||
| 481 | (*Replace outermost quantified variable by Var of given index. | |
| 482 | Could clash with Vars already present.*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 483 | fun forall_elim_var i th = | 
| 0 | 484 |     let val {prop,sign,...} = rep_thm th
 | 
| 485 | in case prop of | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 486 |           Const("all",_) $ Abs(a,T,_) =>
 | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 487 | forall_elim (cterm_of sign (Var((a,i), T))) th | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 488 |         | _ => raise THM("forall_elim_var", i, [th])
 | 
| 0 | 489 | end; | 
| 490 | ||
| 491 | (*Repeat forall_elim_var until all outer quantifiers are removed*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 492 | fun forall_elim_vars i th = | 
| 0 | 493 | forall_elim_vars i (forall_elim_var i th) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 494 | handle THM _ => th; | 
| 0 | 495 | |
| 496 | (*Specialization over a list of cterms*) | |
| 497 | fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th); | |
| 498 | ||
| 499 | (* maps [A1,...,An], B to [| A1;...;An |] ==> B *) | |
| 500 | fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th); | |
| 501 | ||
| 502 | (* maps [| A1;...;An |] ==> B and [A1,...,An] to B *) | |
| 503 | fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths); | |
| 504 | ||
| 505 | (*Reset Var indexes to zero, renaming to preserve distinctness*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 506 | fun zero_var_indexes th = | 
| 0 | 507 |     let val {prop,sign,...} = rep_thm th;
 | 
| 508 | val vars = term_vars prop | |
| 509 | val bs = foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 510 | val inrs = add_term_tvars(prop,[]); | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 511 | val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs)); | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 512 | val tye = map (fn ((v,rs),a) => (v, TVar((a,0),rs))) (inrs ~~ nms') | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 513 | val ctye = map (fn (v,T) => (v,ctyp_of sign T)) tye; | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 514 | fun varpairs([],[]) = [] | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 515 | | varpairs((var as Var(v,T)) :: vars, b::bs) = | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 516 | let val T' = typ_subst_TVars tye T | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 517 | in (cterm_of sign (Var(v,T')), | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 518 | cterm_of sign (Var((b,0),T'))) :: varpairs(vars,bs) | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 519 | end | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 520 |           | varpairs _ = raise TERM("varpairs", []);
 | 
| 0 | 521 | in instantiate (ctye, varpairs(vars,rev bs)) th end; | 
| 522 | ||
| 523 | ||
| 524 | (*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers; | |
| 525 | all generality expressed by Vars having index 0.*) | |
| 526 | fun standard th = | |
| 527 |     let val {maxidx,...} = rep_thm th
 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 528 | in varifyT (zero_var_indexes (forall_elim_vars(maxidx+1) | 
| 0 | 529 | (forall_intr_frees(implies_intr_hyps th)))) | 
| 530 | end; | |
| 531 | ||
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 532 | (*Assume a new formula, read following the same conventions as axioms. | 
| 0 | 533 | Generalizes over Free variables, | 
| 534 | creates the assumption, and then strips quantifiers. | |
| 535 | Example is [| ALL x:?A. ?P(x) |] ==> [| ?P(?a) |] | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 536 | [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ] *) | 
| 0 | 537 | fun assume_ax thy sP = | 
| 538 | let val sign = sign_of thy | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 539 | val prop = Logic.close_form (term_of (read_cterm sign | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 540 | (sP, propT))) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 541 | in forall_elim_vars 0 (assume (cterm_of sign prop)) end; | 
| 0 | 542 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 543 | (*Resolution: exactly one resolvent must be produced.*) | 
| 0 | 544 | fun tha RSN (i,thb) = | 
| 545 | case Sequence.chop (2, biresolution false [(false,tha)] i thb) of | |
| 546 | ([th],_) => th | |
| 547 |     | ([],_)   => raise THM("RSN: no unifiers", i, [tha,thb])
 | |
| 548 |     |      _   => raise THM("RSN: multiple unifiers", i, [tha,thb]);
 | |
| 549 | ||
| 550 | (*resolution: P==>Q, Q==>R gives P==>R. *) | |
| 551 | fun tha RS thb = tha RSN (1,thb); | |
| 552 | ||
| 553 | (*For joining lists of rules*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 554 | fun thas RLN (i,thbs) = | 
| 0 | 555 | let val resolve = biresolution false (map (pair false) thas) i | 
| 556 | fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => [] | |
| 557 | in flat (map resb thbs) end; | |
| 558 | ||
| 559 | fun thas RL thbs = thas RLN (1,thbs); | |
| 560 | ||
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 561 | (*Resolve a list of rules against bottom_rl from right to left; | 
| 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 562 | makes proof trees*) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 563 | fun rls MRS bottom_rl = | 
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 564 | let fun rs_aux i [] = bottom_rl | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 565 | | rs_aux i (rl::rls) = rl RSN (i, rs_aux (i+1) rls) | 
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 566 | in rs_aux 1 rls end; | 
| 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 567 | |
| 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 568 | (*As above, but for rule lists*) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 569 | fun rlss MRL bottom_rls = | 
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 570 | let fun rs_aux i [] = bottom_rls | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 571 | | rs_aux i (rls::rlss) = rls RLN (i, rs_aux (i+1) rlss) | 
| 11 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 572 | in rs_aux 1 rlss end; | 
| 
d0e17c42dbb4
Added MRS, MRL from ZF/ROOT.ML.  These support forward proof, resolving a
 lcp parents: 
0diff
changeset | 573 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 574 | (*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R | 
| 0 | 575 | with no lifting or renaming! Q may contain ==> or meta-quants | 
| 576 | ALWAYS deletes premise i *) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 577 | fun compose(tha,i,thb) = | 
| 0 | 578 | Sequence.list_of_s (bicompose false (false,tha,0) i thb); | 
| 579 | ||
| 580 | (*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*) | |
| 581 | fun tha COMP thb = | |
| 582 | case compose(tha,1,thb) of | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 583 | [th] => th | 
| 0 | 584 |       | _ =>   raise THM("COMP", 1, [tha,thb]);
 | 
| 585 | ||
| 586 | (*Instantiate theorem th, reading instantiations under signature sg*) | |
| 587 | fun read_instantiate_sg sg sinsts th = | |
| 588 | let val ts = types_sorts th; | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 589 | in instantiate (read_insts sg ts ts sinsts) th end; | 
| 0 | 590 | |
| 591 | (*Instantiate theorem th, reading instantiations under theory of th*) | |
| 592 | fun read_instantiate sinsts th = | |
| 593 | read_instantiate_sg (#sign (rep_thm th)) sinsts th; | |
| 594 | ||
| 595 | ||
| 596 | (*Left-to-right replacements: tpairs = [...,(vi,ti),...]. | |
| 597 | Instantiates distinct Vars by terms, inferring type instantiations. *) | |
| 598 | local | |
| 599 | fun add_types ((ct,cu), (sign,tye)) = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 600 |     let val {sign=signt, t=t, T= T, ...} = rep_cterm ct
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 601 |         and {sign=signu, t=u, T= U, ...} = rep_cterm cu
 | 
| 0 | 602 | val sign' = Sign.merge(sign, Sign.merge(signt, signu)) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 603 | val tye' = Type.unify (#tsig(Sign.rep_sg sign')) ((T,U), tye) | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 604 |           handle Type.TUNIFY => raise TYPE("add_types", [T,U], [t,u])
 | 
| 0 | 605 | in (sign', tye') end; | 
| 606 | in | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 607 | fun cterm_instantiate ctpairs0 th = | 
| 0 | 608 | let val (sign,tye) = foldr add_types (ctpairs0, (#sign(rep_thm th),[])) | 
| 609 | val tsig = #tsig(Sign.rep_sg sign); | |
| 610 | fun instT(ct,cu) = let val inst = subst_TVars tye | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 611 | in (cterm_fun inst ct, cterm_fun inst cu) end | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 612 | fun ctyp2 (ix,T) = (ix, ctyp_of sign T) | 
| 0 | 613 | in instantiate (map ctyp2 tye, map instT ctpairs0) th end | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 614 | handle TERM _ => | 
| 0 | 615 |            raise THM("cterm_instantiate: incompatible signatures",0,[th])
 | 
| 616 |        | TYPE _ => raise THM("cterm_instantiate: types", 0, [th])
 | |
| 617 | end; | |
| 618 | ||
| 619 | ||
| 620 | (** theorem equality test is exported and used by BEST_FIRST **) | |
| 621 | ||
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 622 | (*equality of theorems uses equality of signatures and | 
| 0 | 623 | the a-convertible test for terms*) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 624 | fun eq_thm (th1,th2) = | 
| 0 | 625 |     let val {sign=sg1, hyps=hyps1, prop=prop1, ...} = rep_thm th1
 | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 626 |         and {sign=sg2, hyps=hyps2, prop=prop2, ...} = rep_thm th2
 | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 627 | in Sign.eq_sg (sg1,sg2) andalso | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 628 | aconvs(hyps1,hyps2) andalso | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 629 | prop1 aconv prop2 | 
| 0 | 630 | end; | 
| 631 | ||
| 632 | (*Do the two theorems have the same signature?*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 633 | fun eq_thm_sg (th1,th2) = Sign.eq_sg(#sign(rep_thm th1), #sign(rep_thm th2)); | 
| 0 | 634 | |
| 635 | (*Useful "distance" function for BEST_FIRST*) | |
| 636 | val size_of_thm = size_of_term o #prop o rep_thm; | |
| 637 | ||
| 638 | ||
| 639 | (*** Meta-Rewriting Rules ***) | |
| 640 | ||
| 641 | ||
| 642 | val reflexive_thm = | |
| 385 | 643 |   let val cx = cterm_of Sign.pure (Var(("x",0),TVar(("'a",0),logicS)))
 | 
| 0 | 644 | in Thm.reflexive cx end; | 
| 645 | ||
| 646 | val symmetric_thm = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 647 |   let val xy = read_cterm Sign.pure ("x::'a::logic == y",propT)
 | 
| 0 | 648 | in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end; | 
| 649 | ||
| 650 | val transitive_thm = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 651 |   let val xy = read_cterm Sign.pure ("x::'a::logic == y",propT)
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 652 |       val yz = read_cterm Sign.pure ("y::'a::logic == z",propT)
 | 
| 0 | 653 | val xythm = Thm.assume xy and yzthm = Thm.assume yz | 
| 654 | in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end; | |
| 655 | ||
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 656 | (** Below, a "conversion" has type cterm -> thm **) | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 657 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 658 | val refl_cimplies = reflexive (cterm_of Sign.pure implies); | 
| 0 | 659 | |
| 660 | (*In [A1,...,An]==>B, rewrite the selected A's only -- for rewrite_goals_tac*) | |
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 661 | (*Do not rewrite flex-flex pairs*) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 662 | fun goals_conv pred cv = | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 663 | let fun gconv i ct = | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 664 | let val (A,B) = Thm.dest_cimplies ct | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 665 | val (thA,j) = case term_of A of | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 666 |                   Const("=?=",_)$_$_ => (reflexive A, i)
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 667 | | _ => (if pred i then cv A else reflexive A, i+1) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 668 | in combination (combination refl_cimplies thA) (gconv j B) end | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 669 | handle TERM _ => reflexive ct | 
| 0 | 670 | in gconv 1 end; | 
| 671 | ||
| 672 | (*Use a conversion to transform a theorem*) | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 673 | fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th; | 
| 0 | 674 | |
| 675 | (*rewriting conversion*) | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 676 | fun rew_conv mode prover mss = rewrite_cterm mode mss prover; | 
| 0 | 677 | |
| 678 | (*Rewrite a theorem*) | |
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 679 | fun rewrite_rule thms = | 
| 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 680 | fconv_rule (rew_conv (true,false) (K(K None)) (Thm.mss_of thms)); | 
| 0 | 681 | |
| 682 | (*Rewrite the subgoals of a proof state (represented by a theorem) *) | |
| 683 | fun rewrite_goals_rule thms = | |
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 684 | fconv_rule (goals_conv (K true) (rew_conv (true,false) (K(K None)) | 
| 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 685 | (Thm.mss_of thms))); | 
| 0 | 686 | |
| 687 | (*Rewrite the subgoal of a proof state (represented by a theorem) *) | |
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 688 | fun rewrite_goal_rule mode prover mss i thm = | 
| 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 689 | if 0 < i andalso i <= nprems_of thm | 
| 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 690 | then fconv_rule (goals_conv (fn j => j=i) (rew_conv mode prover mss)) thm | 
| 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
211diff
changeset | 691 |   else raise THM("rewrite_goal_rule",i,[thm]);
 | 
| 0 | 692 | |
| 693 | ||
| 694 | (** Derived rules mainly for METAHYPS **) | |
| 695 | ||
| 696 | (*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*) | |
| 697 | fun equal_abs_elim ca eqth = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 698 |   let val {sign=signa, t=a, ...} = rep_cterm ca
 | 
| 0 | 699 | and combth = combination eqth (reflexive ca) | 
| 700 |       val {sign,prop,...} = rep_thm eqth
 | |
| 701 | val (abst,absu) = Logic.dest_equals prop | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 702 | val cterm = cterm_of (Sign.merge (sign,signa)) | 
| 0 | 703 | in transitive (symmetric (beta_conversion (cterm (abst$a)))) | 
| 704 | (transitive combth (beta_conversion (cterm (absu$a)))) | |
| 705 | end | |
| 706 |   handle THM _ => raise THM("equal_abs_elim", 0, [eqth]);
 | |
| 707 | ||
| 708 | (*Calling equal_abs_elim with multiple terms*) | |
| 709 | fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th); | |
| 710 | ||
| 711 | local | |
| 712 | open Logic | |
| 713 |   val alpha = TVar(("'a",0), [])     (*  type ?'a::{}  *)
 | |
| 714 |   fun err th = raise THM("flexpair_inst: ", 0, [th])
 | |
| 715 | fun flexpair_inst def th = | |
| 716 |     let val {prop = Const _ $ t $ u,  sign,...} = rep_thm th
 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 717 | val cterm = cterm_of sign | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 718 | fun cvar a = cterm(Var((a,0),alpha)) | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 719 | val def' = cterm_instantiate [(cvar"t", cterm t), (cvar"u", cterm u)] | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 720 | def | 
| 0 | 721 | in equal_elim def' th | 
| 722 | end | |
| 723 | handle THM _ => err th | bind => err th | |
| 724 | in | |
| 725 | val flexpair_intr = flexpair_inst (symmetric flexpair_def) | |
| 726 | and flexpair_elim = flexpair_inst flexpair_def | |
| 727 | end; | |
| 728 | ||
| 729 | (*Version for flexflex pairs -- this supports lifting.*) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 730 | fun flexpair_abs_elim_list cts = | 
| 0 | 731 | flexpair_intr o equal_abs_elim_list cts o flexpair_elim; | 
| 732 | ||
| 733 | ||
| 734 | (*** Some useful meta-theorems ***) | |
| 735 | ||
| 736 | (*The rule V/V, obtains assumption solving for eresolve_tac*) | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 737 | val asm_rl = trivial(read_cterm Sign.pure ("PROP ?psi",propT));
 | 
| 0 | 738 | |
| 739 | (*Meta-level cut rule: [| V==>W; V |] ==> W *) | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 740 | val cut_rl = trivial(read_cterm Sign.pure | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 741 |         ("PROP ?psi ==> PROP ?theta", propT));
 | 
| 0 | 742 | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 743 | (*Generalized elim rule for one conclusion; cut_rl with reversed premises: | 
| 0 | 744 | [| PROP V; PROP V ==> PROP W |] ==> PROP W *) | 
| 745 | val revcut_rl = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 746 |   let val V = read_cterm Sign.pure ("PROP V", propT)
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 747 |       and VW = read_cterm Sign.pure ("PROP V ==> PROP W", propT);
 | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 748 | in standard (implies_intr V | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 749 | (implies_intr VW | 
| 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 750 | (implies_elim (assume VW) (assume V)))) | 
| 0 | 751 | end; | 
| 752 | ||
| 668 | 753 | (*for deleting an unwanted assumption*) | 
| 754 | val thin_rl = | |
| 755 |   let val V = read_cterm Sign.pure ("PROP V", propT)
 | |
| 756 |       and W = read_cterm Sign.pure ("PROP W", propT);
 | |
| 757 | in standard (implies_intr V (implies_intr W (assume W))) | |
| 758 | end; | |
| 759 | ||
| 0 | 760 | (* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*) | 
| 761 | val triv_forall_equality = | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 762 |   let val V  = read_cterm Sign.pure ("PROP V", propT)
 | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
214diff
changeset | 763 |       and QV = read_cterm Sign.pure ("!!x::'a. PROP V", propT)
 | 
| 385 | 764 |       and x  = read_cterm Sign.pure ("x", TFree("'a",logicS));
 | 
| 0 | 765 | in standard (equal_intr (implies_intr QV (forall_elim x (assume QV))) | 
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 766 | (implies_intr V (forall_intr x (assume V)))) | 
| 0 | 767 | end; | 
| 768 | ||
| 769 | end | |
| 770 | end; | |
| 252 
7532f95d7f44
removed eq_sg, pprint_sg, print_sg (now in sign.ML);
 wenzelm parents: 
229diff
changeset | 771 |