| author | wenzelm | 
| Wed, 29 Apr 1998 11:24:58 +0200 | |
| changeset 4850 | 050481f41e28 | 
| parent 4847 | ea7d7a65e4e9 | 
| child 4999 | 4c74267cfa0c | 
| permissions | -rw-r--r-- | 
| 250 | 1 | (* Title: Pure/thm.ML | 
| 0 | 2 | ID: $Id$ | 
| 250 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 4 | Copyright 1994 University of Cambridge | 
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 5 | |
| 1160 | 6 | The core of Isabelle's Meta Logic: certified types and terms, meta | 
| 1529 | 7 | theorems, meta rules (including resolution and simplification). | 
| 0 | 8 | *) | 
| 9 | ||
| 250 | 10 | signature THM = | 
| 1503 | 11 | sig | 
| 1160 | 12 | (*certified types*) | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 13 | type ctyp | 
| 1238 | 14 |   val rep_ctyp          : ctyp -> {sign: Sign.sg, T: typ}
 | 
| 15 | val typ_of : ctyp -> typ | |
| 16 | val ctyp_of : Sign.sg -> typ -> ctyp | |
| 17 | val read_ctyp : Sign.sg -> string -> ctyp | |
| 1160 | 18 | |
| 19 | (*certified terms*) | |
| 20 | type cterm | |
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 21 | exception CTERM of string | 
| 4270 | 22 |   val rep_cterm         : cterm -> {sign: Sign.sg, t: term, T: typ, maxidx: int}
 | 
| 4288 | 23 |   val crep_cterm        : cterm -> {sign: Sign.sg, t: term, T: ctyp, maxidx: int}
 | 
| 1238 | 24 | val term_of : cterm -> term | 
| 25 | val cterm_of : Sign.sg -> term -> cterm | |
| 2671 | 26 | val ctyp_of_term : cterm -> ctyp | 
| 1238 | 27 | val read_cterm : Sign.sg -> string * typ -> cterm | 
| 28 | val cterm_fun : (term -> term) -> (cterm -> cterm) | |
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 29 | val dest_comb : cterm -> cterm * cterm | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 30 | val dest_abs : cterm -> cterm * cterm | 
| 1703 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 clasohm parents: 
1659diff
changeset | 31 | val adjust_maxidx : cterm -> cterm | 
| 1516 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 clasohm parents: 
1503diff
changeset | 32 | val capply : cterm -> cterm -> cterm | 
| 1517 | 33 | val cabs : cterm -> cterm -> cterm | 
| 1238 | 34 | val read_def_cterm : | 
| 1160 | 35 | Sign.sg * (indexname -> typ option) * (indexname -> sort option) -> | 
| 36 | string list -> bool -> string * typ -> cterm * (indexname * typ) list | |
| 4281 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 37 | val read_def_cterms : | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 38 | Sign.sg * (indexname -> typ option) * (indexname -> sort option) -> | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 39 | string list -> bool -> (string * typ)list | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 40 | -> cterm list * (indexname * typ)list | 
| 1160 | 41 | |
| 2671 | 42 | (*proof terms [must DUPLICATE declaration as a specification]*) | 
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 43 | datatype deriv_kind = MinDeriv | ThmDeriv | FullDeriv; | 
| 2386 | 44 | val keep_derivs : deriv_kind ref | 
| 1529 | 45 | datatype rule = | 
| 2386 | 46 | MinProof | 
| 4182 | 47 | | Oracle of string * Sign.sg * object | 
| 48 | | Axiom of string | |
| 2671 | 49 | | Theorem of string | 
| 50 | | Assume of cterm | |
| 51 | | Implies_intr of cterm | |
| 1529 | 52 | | Implies_intr_shyps | 
| 53 | | Implies_intr_hyps | |
| 54 | | Implies_elim | |
| 2671 | 55 | | Forall_intr of cterm | 
| 56 | | Forall_elim of cterm | |
| 57 | | Reflexive of cterm | |
| 1529 | 58 | | Symmetric | 
| 59 | | Transitive | |
| 2671 | 60 | | Beta_conversion of cterm | 
| 1529 | 61 | | Extensional | 
| 2671 | 62 | | Abstract_rule of string * cterm | 
| 1529 | 63 | | Combination | 
| 64 | | Equal_intr | |
| 65 | | Equal_elim | |
| 2671 | 66 | | Trivial of cterm | 
| 67 | | Lift_rule of cterm * int | |
| 68 | | Assumption of int * Envir.env option | |
| 69 | | Rotate_rule of int * int | |
| 70 | | Instantiate of (indexname * ctyp) list * (cterm * cterm) list | |
| 71 | | Bicompose of bool * bool * int * int * Envir.env | |
| 72 | | Flexflex_rule of Envir.env | |
| 4182 | 73 | | Class_triv of class | 
| 1529 | 74 | | VarifyT | 
| 75 | | FreezeT | |
| 2671 | 76 | | RewriteC of cterm | 
| 77 | | CongC of cterm | |
| 78 | | Rewrite_cterm of cterm | |
| 79 | | Rename_params_rule of string list * int; | |
| 1529 | 80 | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 81 | type deriv (* = rule mtree *) | 
| 1529 | 82 | |
| 1160 | 83 | (*meta theorems*) | 
| 84 | type thm | |
| 85 | exception THM of string * int * thm list | |
| 1529 | 86 |   val rep_thm           : thm -> {sign: Sign.sg, der: deriv, maxidx: int,
 | 
| 2386 | 87 | shyps: sort list, hyps: term list, | 
| 88 | prop: term} | |
| 1529 | 89 |   val crep_thm          : thm -> {sign: Sign.sg, der: deriv, maxidx: int,
 | 
| 2386 | 90 | shyps: sort list, hyps: cterm list, | 
| 91 | prop: cterm} | |
| 3994 | 92 | val eq_thm : thm * thm -> bool | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 93 | val sign_of_thm : thm -> Sign.sg | 
| 4254 | 94 | val transfer_sg : Sign.sg -> thm -> thm | 
| 3895 | 95 | val transfer : theory -> thm -> thm | 
| 1238 | 96 | val tpairs_of : thm -> (term * term) list | 
| 97 | val prems_of : thm -> term list | |
| 98 | val nprems_of : thm -> int | |
| 99 | val concl_of : thm -> term | |
| 100 | val cprop_of : thm -> cterm | |
| 101 | val extra_shyps : thm -> sort list | |
| 3061 | 102 | val force_strip_shyps : bool ref (* FIXME tmp (since 1995/08/01) *) | 
| 1238 | 103 | val strip_shyps : thm -> thm | 
| 104 | val implies_intr_shyps: thm -> thm | |
| 3812 | 105 | val get_axiom : theory -> xstring -> thm | 
| 4847 | 106 | val get_def : theory -> xstring -> thm | 
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 107 | val name_thm : string * thm -> thm | 
| 4018 | 108 | val name_of_thm : thm -> string | 
| 1238 | 109 | val axioms_of : theory -> (string * thm) list | 
| 1160 | 110 | |
| 111 | (*meta rules*) | |
| 1238 | 112 | val assume : cterm -> thm | 
| 1416 | 113 | val compress : thm -> thm | 
| 1238 | 114 | val implies_intr : cterm -> thm -> thm | 
| 115 | val implies_elim : thm -> thm -> thm | |
| 116 | val forall_intr : cterm -> thm -> thm | |
| 117 | val forall_elim : cterm -> thm -> thm | |
| 118 | val reflexive : cterm -> thm | |
| 119 | val symmetric : thm -> thm | |
| 120 | val transitive : thm -> thm -> thm | |
| 121 | val beta_conversion : cterm -> thm | |
| 122 | val extensional : thm -> thm | |
| 123 | val abstract_rule : string -> cterm -> thm -> thm | |
| 124 | val combination : thm -> thm -> thm | |
| 125 | val equal_intr : thm -> thm -> thm | |
| 126 | val equal_elim : thm -> thm -> thm | |
| 127 | val implies_intr_hyps : thm -> thm | |
| 4270 | 128 | val flexflex_rule : thm -> thm Seq.seq | 
| 1238 | 129 | val instantiate : | 
| 1160 | 130 | (indexname * ctyp) list * (cterm * cterm) list -> thm -> thm | 
| 1238 | 131 | val trivial : cterm -> thm | 
| 132 | val class_triv : theory -> class -> thm | |
| 133 | val varifyT : thm -> thm | |
| 134 | val freezeT : thm -> thm | |
| 135 | val dest_state : thm * int -> | |
| 1160 | 136 | (term * term) list * term list * term * term | 
| 1238 | 137 | val lift_rule : (thm * int) -> thm -> thm | 
| 4270 | 138 | val assumption : int -> thm -> thm Seq.seq | 
| 1238 | 139 | val eq_assumption : int -> thm -> thm | 
| 2671 | 140 | val rotate_rule : int -> int -> thm -> thm | 
| 1160 | 141 | val rename_params_rule: string list * int -> thm -> thm | 
| 1238 | 142 | val bicompose : bool -> bool * thm * int -> | 
| 4270 | 143 | int -> thm -> thm Seq.seq | 
| 1238 | 144 | val biresolution : bool -> (bool * thm) list -> | 
| 4270 | 145 | int -> thm -> thm Seq.seq | 
| 1160 | 146 | |
| 147 | (*meta simplification*) | |
| 3550 | 148 | exception SIMPLIFIER of string * thm | 
| 1160 | 149 | type meta_simpset | 
| 3550 | 150 | val dest_mss : meta_simpset -> | 
| 151 |     {simps: thm list, congs: thm list, procs: (string * cterm list) list}
 | |
| 1238 | 152 | val empty_mss : meta_simpset | 
| 3550 | 153 | val merge_mss : meta_simpset * meta_simpset -> meta_simpset | 
| 1238 | 154 | val add_simps : meta_simpset * thm list -> meta_simpset | 
| 155 | val del_simps : meta_simpset * thm list -> meta_simpset | |
| 156 | val mss_of : thm list -> meta_simpset | |
| 157 | val add_congs : meta_simpset * thm list -> meta_simpset | |
| 2626 | 158 | val del_congs : meta_simpset * thm list -> meta_simpset | 
| 2509 | 159 | val add_simprocs : meta_simpset * | 
| 3577 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 160 | (string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list | 
| 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 161 | -> meta_simpset | 
| 2509 | 162 | val del_simprocs : meta_simpset * | 
| 3577 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 163 | (string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp) list | 
| 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 164 | -> meta_simpset | 
| 1238 | 165 | val add_prems : meta_simpset * thm list -> meta_simpset | 
| 166 | val prems_of_mss : meta_simpset -> thm list | |
| 167 | val set_mk_rews : meta_simpset * (thm -> thm list) -> meta_simpset | |
| 4679 | 168 | val set_mk_sym : meta_simpset * (thm -> thm option) -> meta_simpset | 
| 169 | val set_mk_eq_True : meta_simpset * (thm -> thm option) -> meta_simpset | |
| 2509 | 170 | val set_termless : meta_simpset * (term * term -> bool) -> meta_simpset | 
| 1238 | 171 | val trace_simp : bool ref | 
| 4713 | 172 | val rewrite_cterm : bool * bool * bool -> meta_simpset -> | 
| 1529 | 173 | (meta_simpset -> thm -> thm option) -> cterm -> thm | 
| 1539 | 174 | |
| 4124 | 175 | val invoke_oracle : theory -> xstring -> Sign.sg * object -> thm | 
| 250 | 176 | end; | 
| 0 | 177 | |
| 3550 | 178 | structure Thm: THM = | 
| 0 | 179 | struct | 
| 250 | 180 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 181 | (*** Certified terms and types ***) | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 182 | |
| 250 | 183 | (** certified types **) | 
| 184 | ||
| 185 | (*certified typs under a signature*) | |
| 186 | ||
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 187 | datatype ctyp = Ctyp of {sign_ref: Sign.sg_ref, T: typ};
 | 
| 250 | 188 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 189 | fun rep_ctyp (Ctyp {sign_ref, T}) = {sign = Sign.deref sign_ref, T = T};
 | 
| 250 | 190 | fun typ_of (Ctyp {T, ...}) = T;
 | 
| 191 | ||
| 192 | fun ctyp_of sign T = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 193 |   Ctyp {sign_ref = Sign.self_ref sign, T = Sign.certify_typ sign T};
 | 
| 250 | 194 | |
| 195 | fun read_ctyp sign s = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 196 |   Ctyp {sign_ref = Sign.self_ref sign, T = Sign.read_typ (sign, K None) s};
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 197 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 198 | |
| 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 199 | |
| 250 | 200 | (** certified terms **) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 201 | |
| 250 | 202 | (*certified terms under a signature, with checked typ and maxidx of Vars*) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 203 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 204 | datatype cterm = Cterm of {sign_ref: Sign.sg_ref, t: term, T: typ, maxidx: int};
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 205 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 206 | fun rep_cterm (Cterm {sign_ref, t, T, maxidx}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 207 |   {sign = Sign.deref sign_ref, t = t, T = T, maxidx = maxidx};
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 208 | |
| 4288 | 209 | fun crep_cterm (Cterm {sign_ref, t, T, maxidx}) =
 | 
| 210 |   {sign = Sign.deref sign_ref, t = t, T = Ctyp {sign_ref = sign_ref, T = T},
 | |
| 211 | maxidx = maxidx}; | |
| 212 | ||
| 250 | 213 | fun term_of (Cterm {t, ...}) = t;
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 214 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 215 | fun ctyp_of_term (Cterm {sign_ref, T, ...}) = Ctyp {sign_ref = sign_ref, T = T};
 | 
| 2671 | 216 | |
| 250 | 217 | (*create a cterm by checking a "raw" term with respect to a signature*) | 
| 218 | fun cterm_of sign tm = | |
| 219 | let val (t, T, maxidx) = Sign.certify_term sign tm | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 220 |   in  Cterm {sign_ref = Sign.self_ref sign, t = t, T = T, maxidx = maxidx}
 | 
| 1394 
a1d2735f5ade
New function read_cterms is a combination of read_def_cterm and
 paulson parents: 
1258diff
changeset | 221 | end; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 222 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 223 | fun cterm_fun f (Cterm {sign_ref, t, ...}) = cterm_of (Sign.deref sign_ref) (f t);
 | 
| 250 | 224 | |
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 225 | |
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 226 | exception CTERM of string; | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 227 | |
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 228 | (*Destruct application in cterms*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 229 | fun dest_comb (Cterm {sign_ref, T, maxidx, t = A $ B}) =
 | 
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 230 | let val typeA = fastype_of A; | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 231 | val typeB = | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 232 |             case typeA of Type("fun",[S,T]) => S
 | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 233 | | _ => error "Function type expected in dest_comb"; | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 234 | in | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 235 |       (Cterm {sign_ref=sign_ref, maxidx=maxidx, t=A, T=typeA},
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 236 |        Cterm {sign_ref=sign_ref, maxidx=maxidx, t=B, T=typeB})
 | 
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 237 | end | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 238 | | dest_comb _ = raise CTERM "dest_comb"; | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 239 | |
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 240 | (*Destruct abstraction in cterms*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 241 | fun dest_abs (Cterm {sign_ref, T as Type("fun",[_,S]), maxidx, t=Abs(x,ty,M)}) = 
 | 
| 1516 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 clasohm parents: 
1503diff
changeset | 242 | let val (y,N) = variant_abs (x,ty,M) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 243 |       in (Cterm {sign_ref = sign_ref, T = ty, maxidx = 0, t = Free(y,ty)},
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 244 |           Cterm {sign_ref = sign_ref, T = S, maxidx = maxidx, t = N})
 | 
| 1493 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 245 | end | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 246 | | dest_abs _ = raise CTERM "dest_abs"; | 
| 
e936723cb94d
added dest_comb, dest_abs and mk_prop for manipulating cterms
 clasohm parents: 
1460diff
changeset | 247 | |
| 2147 | 248 | (*Makes maxidx precise: it is often too big*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 249 | fun adjust_maxidx (ct as Cterm {sign_ref, T, t, maxidx, ...}) =
 | 
| 2147 | 250 | if maxidx = ~1 then ct | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 251 |   else  Cterm {sign_ref = sign_ref, T = T, maxidx = maxidx_of_term t, t = t};
 | 
| 1703 
e22ad43bab5f
moved dest_cimplies to drule.ML; added adjust_maxidx
 clasohm parents: 
1659diff
changeset | 252 | |
| 1516 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 clasohm parents: 
1503diff
changeset | 253 | (*Form cterm out of a function and an argument*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 254 | fun capply (Cterm {t=f, sign_ref=sign_ref1, T=Type("fun",[dty,rty]), maxidx=maxidx1})
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 255 |            (Cterm {t=x, sign_ref=sign_ref2, T, maxidx=maxidx2}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 256 |       if T = dty then Cterm{t=f$x, sign_ref=Sign.merge_refs(sign_ref1,sign_ref2), T=rty,
 | 
| 2147 | 257 | maxidx=Int.max(maxidx1, maxidx2)} | 
| 1516 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 clasohm parents: 
1503diff
changeset | 258 | else raise CTERM "capply: types don't agree" | 
| 
96286c4e32de
removed mk_prop; added capply; simplified dest_abs
 clasohm parents: 
1503diff
changeset | 259 | | capply _ _ = raise CTERM "capply: first arg is not a function" | 
| 250 | 260 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 261 | fun cabs (Cterm {t=Free(a,ty), sign_ref=sign_ref1, T=T1, maxidx=maxidx1})
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 262 |          (Cterm {t=t2, sign_ref=sign_ref2, T=T2, maxidx=maxidx2}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 263 |       Cterm {t=absfree(a,ty,t2), sign_ref=Sign.merge_refs(sign_ref1,sign_ref2),
 | 
| 2147 | 264 | T = ty --> T2, maxidx=Int.max(maxidx1, maxidx2)} | 
| 1517 | 265 | | cabs _ _ = raise CTERM "cabs: first arg is not a free variable"; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 266 | |
| 2509 | 267 | |
| 268 | ||
| 574 | 269 | (** read cterms **) (*exception ERROR*) | 
| 250 | 270 | |
| 4281 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 271 | (*read terms, infer types, certify terms*) | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 272 | fun read_def_cterms (sign, types, sorts) used freeze sTs = | 
| 250 | 273 | let | 
| 4281 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 274 | val syn = #syn (Sign.rep_sg sign) | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 275 | fun read(s,T) = | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 276 | let val T' = Sign.certify_typ sign T | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 277 | handle TYPE (msg, _, _) => error msg | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 278 | in (Syntax.read syn T' s, T') end | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 279 | val tsTs = map read sTs | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 280 | val (ts',tye) = Sign.infer_types_simult sign types sorts used freeze tsTs; | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 281 | val cts = map (cterm_of sign) ts' | 
| 2979 | 282 | handle TYPE (msg, _, _) => error msg | 
| 2386 | 283 | | TERM (msg, _) => error msg; | 
| 4281 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 284 | in (cts, tye) end; | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 285 | |
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 286 | (*read term, infer types, certify term*) | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 287 | fun read_def_cterm args used freeze aT = | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 288 | let val ([ct],tye) = read_def_cterms args used freeze [aT] | 
| 
6c6073b13600
Added read_def_cterms for simultaneous reading/typing of terms under
 nipkow parents: 
4270diff
changeset | 289 | in (ct,tye) end; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 290 | |
| 949 
83c588d6fee9
Changed treatment of during type inference internally generated type
 nipkow parents: 
922diff
changeset | 291 | fun read_cterm sign = #1 o read_def_cterm (sign, K None, K None) [] true; | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 292 | |
| 250 | 293 | |
| 294 | ||
| 1529 | 295 | (*** Derivations ***) | 
| 296 | ||
| 297 | (*Names of rules in derivations. Includes logically trivial rules, if | |
| 298 | executed in ML.*) | |
| 299 | datatype rule = | |
| 2386 | 300 | MinProof (*for building minimal proof terms*) | 
| 4182 | 301 | | Oracle of string * Sign.sg * object (*oracles*) | 
| 1529 | 302 | (*Axioms/theorems*) | 
| 4182 | 303 | | Axiom of string | 
| 2386 | 304 | | Theorem of string | 
| 1529 | 305 | (*primitive inferences and compound versions of them*) | 
| 2386 | 306 | | Assume of cterm | 
| 307 | | Implies_intr of cterm | |
| 1529 | 308 | | Implies_intr_shyps | 
| 309 | | Implies_intr_hyps | |
| 310 | | Implies_elim | |
| 2386 | 311 | | Forall_intr of cterm | 
| 312 | | Forall_elim of cterm | |
| 313 | | Reflexive of cterm | |
| 1529 | 314 | | Symmetric | 
| 315 | | Transitive | |
| 2386 | 316 | | Beta_conversion of cterm | 
| 1529 | 317 | | Extensional | 
| 2386 | 318 | | Abstract_rule of string * cterm | 
| 1529 | 319 | | Combination | 
| 320 | | Equal_intr | |
| 321 | | Equal_elim | |
| 322 | (*derived rules for tactical proof*) | |
| 2386 | 323 | | Trivial of cterm | 
| 324 | (*For lift_rule, the proof state is not a premise. | |
| 325 | Use cterm instead of thm to avoid mutual recursion.*) | |
| 326 | | Lift_rule of cterm * int | |
| 327 | | Assumption of int * Envir.env option (*includes eq_assumption*) | |
| 2671 | 328 | | Rotate_rule of int * int | 
| 2386 | 329 | | Instantiate of (indexname * ctyp) list * (cterm * cterm) list | 
| 330 | | Bicompose of bool * bool * int * int * Envir.env | |
| 331 | | Flexflex_rule of Envir.env (*identifies unifier chosen*) | |
| 1529 | 332 | (*other derived rules*) | 
| 4182 | 333 | | Class_triv of class | 
| 1529 | 334 | | VarifyT | 
| 335 | | FreezeT | |
| 336 | (*for the simplifier*) | |
| 2386 | 337 | | RewriteC of cterm | 
| 338 | | CongC of cterm | |
| 339 | | Rewrite_cterm of cterm | |
| 1529 | 340 | (*Logical identities, recorded since they are part of the proof process*) | 
| 2386 | 341 | | Rename_params_rule of string list * int; | 
| 1529 | 342 | |
| 343 | ||
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 344 | type deriv = rule mtree; | 
| 1529 | 345 | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 346 | datatype deriv_kind = MinDeriv | ThmDeriv | FullDeriv; | 
| 1529 | 347 | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 348 | val keep_derivs = ref MinDeriv; | 
| 1529 | 349 | |
| 350 | ||
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 351 | (*Build a minimal derivation. Keep oracles; suppress atomic inferences; | 
| 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 352 | retain Theorems or their underlying links; keep anything else*) | 
| 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 353 | fun squash_derivs [] = [] | 
| 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 354 | | squash_derivs (der::ders) = | 
| 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 355 | (case der of | 
| 2386 | 356 | Join (Oracle _, _) => der :: squash_derivs ders | 
| 357 | | Join (Theorem _, [der']) => if !keep_derivs=ThmDeriv | |
| 358 | then der :: squash_derivs ders | |
| 359 | else squash_derivs (der'::ders) | |
| 360 | | Join (Axiom _, _) => if !keep_derivs=ThmDeriv | |
| 361 | then der :: squash_derivs ders | |
| 362 | else squash_derivs ders | |
| 363 | | Join (_, []) => squash_derivs ders | |
| 364 | | _ => der :: squash_derivs ders); | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 365 | |
| 1529 | 366 | |
| 367 | (*Ensure sharing of the most likely derivation, the empty one!*) | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 368 | val min_infer = Join (MinProof, []); | 
| 1529 | 369 | |
| 370 | (*Make a minimal inference*) | |
| 371 | fun make_min_infer [] = min_infer | |
| 372 | | make_min_infer [der] = der | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 373 | | make_min_infer ders = Join (MinProof, ders); | 
| 1529 | 374 | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 375 | fun infer_derivs (rl, []) = Join (rl, []) | 
| 1529 | 376 | | infer_derivs (rl, ders) = | 
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 377 | if !keep_derivs=FullDeriv then Join (rl, ders) | 
| 1529 | 378 | else make_min_infer (squash_derivs ders); | 
| 379 | ||
| 380 | ||
| 2509 | 381 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 382 | (*** Meta theorems ***) | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 383 | |
| 0 | 384 | datatype thm = Thm of | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 385 |  {sign_ref: Sign.sg_ref,       (*mutable reference to signature*)
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 386 | der: deriv, (*derivation*) | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 387 | maxidx: int, (*maximum index of any Var or TVar*) | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 388 | shyps: sort list, (*sort hypotheses*) | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 389 | hyps: term list, (*hypotheses*) | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 390 | prop: term}; (*conclusion*) | 
| 0 | 391 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 392 | fun rep_thm (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 393 |   {sign = Sign.deref sign_ref, der = der, maxidx = maxidx,
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 394 | shyps = shyps, hyps = hyps, prop = prop}; | 
| 0 | 395 | |
| 1529 | 396 | (*Version of rep_thm returning cterms instead of terms*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 397 | fun crep_thm (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 398 |   let fun ctermf max t = Cterm{sign_ref=sign_ref, t=t, T=propT, maxidx=max};
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 399 |   in {sign = Sign.deref sign_ref, der = der, maxidx = maxidx, shyps = shyps,
 | 
| 1529 | 400 | hyps = map (ctermf ~1) hyps, | 
| 401 | prop = ctermf maxidx prop} | |
| 1517 | 402 | end; | 
| 403 | ||
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 404 | (*errors involving theorems*) | 
| 0 | 405 | exception THM of string * int * thm list; | 
| 406 | ||
| 3994 | 407 | (*equality of theorems uses equality of signatures and the | 
| 408 | a-convertible test for terms*) | |
| 409 | fun eq_thm (th1, th2) = | |
| 410 | let | |
| 411 |     val {sign = sg1, shyps = shyps1, hyps = hyps1, prop = prop1, ...} = rep_thm th1;
 | |
| 412 |     val {sign = sg2, shyps = shyps2, hyps = hyps2, prop = prop2, ...} = rep_thm th2;
 | |
| 413 | in | |
| 414 | Sign.eq_sg (sg1, sg2) andalso | |
| 415 | eq_set_sort (shyps1, shyps2) andalso | |
| 416 | aconvs (hyps1, hyps2) andalso | |
| 417 | prop1 aconv prop2 | |
| 418 | end; | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 419 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 420 | fun sign_of_thm (Thm {sign_ref, ...}) = Sign.deref sign_ref;
 | 
| 0 | 421 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 422 | (*merge signatures of two theorems; raise exception if incompatible*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 423 | fun merge_thm_sgs | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 424 |     (th1 as Thm {sign_ref = sgr1, ...}, th2 as Thm {sign_ref = sgr2, ...}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 425 | Sign.merge_refs (sgr1, sgr2) handle TERM (msg, _) => raise THM (msg, 0, [th1, th2]); | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 426 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 427 | (*transfer thm to super theory (non-destructive)*) | 
| 4254 | 428 | fun transfer_sg sign' thm = | 
| 3895 | 429 | let | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 430 |     val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 431 | val sign = Sign.deref sign_ref; | 
| 3895 | 432 | in | 
| 4254 | 433 | if Sign.eq_sg (sign, sign') then thm | 
| 434 | else if Sign.subsig (sign, sign') then | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 435 |       Thm {sign_ref = Sign.self_ref sign', der = der, maxidx = maxidx,
 | 
| 3895 | 436 | shyps = shyps, hyps = hyps, prop = prop} | 
| 437 |     else raise THM ("transfer: not a super theory", 0, [thm])
 | |
| 438 | end; | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 439 | |
| 4254 | 440 | val transfer = transfer_sg o sign_of; | 
| 441 | ||
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 442 | (*maps object-rule to tpairs*) | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 443 | fun tpairs_of (Thm {prop, ...}) = #1 (Logic.strip_flexpairs prop);
 | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 444 | |
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 445 | (*maps object-rule to premises*) | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 446 | fun prems_of (Thm {prop, ...}) =
 | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 447 | Logic.strip_imp_prems (Logic.skip_flexpairs prop); | 
| 0 | 448 | |
| 449 | (*counts premises in a rule*) | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 450 | fun nprems_of (Thm {prop, ...}) =
 | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 451 | Logic.count_prems (Logic.skip_flexpairs prop, 0); | 
| 0 | 452 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 453 | (*maps object-rule to conclusion*) | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 454 | fun concl_of (Thm {prop, ...}) = Logic.strip_imp_concl prop;
 | 
| 0 | 455 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 456 | (*the statement of any thm is a cterm*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 457 | fun cprop_of (Thm {sign_ref, maxidx, prop, ...}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 458 |   Cterm {sign_ref = sign_ref, maxidx = maxidx, T = propT, t = prop};
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 459 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 460 | |
| 0 | 461 | |
| 1238 | 462 | (** sort contexts of theorems **) | 
| 463 | ||
| 464 | (* basic utils *) | |
| 465 | ||
| 2163 | 466 | (*accumulate sorts suppressing duplicates; these are coded low levelly | 
| 1238 | 467 | to improve efficiency a bit*) | 
| 468 | ||
| 469 | fun add_typ_sorts (Type (_, Ts), Ss) = add_typs_sorts (Ts, Ss) | |
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 470 | | add_typ_sorts (TFree (_, S), Ss) = ins_sort(S,Ss) | 
| 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 471 | | add_typ_sorts (TVar (_, S), Ss) = ins_sort(S,Ss) | 
| 1238 | 472 | and add_typs_sorts ([], Ss) = Ss | 
| 473 | | add_typs_sorts (T :: Ts, Ss) = add_typs_sorts (Ts, add_typ_sorts (T, Ss)); | |
| 474 | ||
| 475 | fun add_term_sorts (Const (_, T), Ss) = add_typ_sorts (T, Ss) | |
| 476 | | add_term_sorts (Free (_, T), Ss) = add_typ_sorts (T, Ss) | |
| 477 | | add_term_sorts (Var (_, T), Ss) = add_typ_sorts (T, Ss) | |
| 478 | | add_term_sorts (Bound _, Ss) = Ss | |
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 479 | | add_term_sorts (Abs (_,T,t), Ss) = add_term_sorts (t, add_typ_sorts (T,Ss)) | 
| 1238 | 480 | | add_term_sorts (t $ u, Ss) = add_term_sorts (t, add_term_sorts (u, Ss)); | 
| 481 | ||
| 482 | fun add_terms_sorts ([], Ss) = Ss | |
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 483 | | add_terms_sorts (t::ts, Ss) = add_terms_sorts (ts, add_term_sorts (t,Ss)); | 
| 1238 | 484 | |
| 1258 | 485 | fun env_codT (Envir.Envir {iTs, ...}) = map snd iTs;
 | 
| 486 | ||
| 487 | fun add_env_sorts (env, Ss) = | |
| 488 | add_terms_sorts (map snd (Envir.alist_of env), | |
| 489 | add_typs_sorts (env_codT env, Ss)); | |
| 490 | ||
| 1238 | 491 | fun add_thm_sorts (Thm {hyps, prop, ...}, Ss) =
 | 
| 492 | add_terms_sorts (hyps, add_term_sorts (prop, Ss)); | |
| 493 | ||
| 494 | fun add_thms_shyps ([], Ss) = Ss | |
| 495 |   | add_thms_shyps (Thm {shyps, ...} :: ths, Ss) =
 | |
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 496 | add_thms_shyps (ths, union_sort(shyps,Ss)); | 
| 1238 | 497 | |
| 498 | ||
| 499 | (*get 'dangling' sort constraints of a thm*) | |
| 500 | fun extra_shyps (th as Thm {shyps, ...}) =
 | |
| 501 | shyps \\ add_thm_sorts (th, []); | |
| 502 | ||
| 503 | ||
| 504 | (* fix_shyps *) | |
| 505 | ||
| 506 | (*preserve sort contexts of rule premises and substituted types*) | |
| 507 | fun fix_shyps thms Ts thm = | |
| 508 | let | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 509 |     val Thm {sign_ref, der, maxidx, hyps, prop, ...} = thm;
 | 
| 1238 | 510 | val shyps = | 
| 511 | add_thm_sorts (thm, add_typs_sorts (Ts, add_thms_shyps (thms, []))); | |
| 512 | in | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 513 |     Thm {sign_ref = sign_ref,
 | 
| 2386 | 514 | der = der, (*No new derivation, as other rules call this*) | 
| 515 | maxidx = maxidx, | |
| 516 | shyps = shyps, hyps = hyps, prop = prop} | |
| 1238 | 517 | end; | 
| 518 | ||
| 519 | ||
| 520 | (* strip_shyps *) (* FIXME improve? (e.g. only minimal extra sorts) *) | |
| 521 | ||
| 3061 | 522 | val force_strip_shyps = ref true; (* FIXME tmp (since 1995/08/01) *) | 
| 1238 | 523 | |
| 524 | (*remove extra sorts that are known to be syntactically non-empty*) | |
| 525 | fun strip_shyps thm = | |
| 526 | let | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 527 |     val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 1238 | 528 | val sorts = add_thm_sorts (thm, []); | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 529 | val maybe_empty = not o Sign.nonempty_sort (Sign.deref sign_ref) sorts; | 
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 530 | val shyps' = filter (fn S => mem_sort(S,sorts) orelse maybe_empty S) shyps; | 
| 1238 | 531 | in | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 532 |     Thm {sign_ref = sign_ref, der = der, maxidx = maxidx,
 | 
| 2386 | 533 | shyps = | 
| 534 | (if eq_set_sort (shyps',sorts) orelse | |
| 535 | not (!force_strip_shyps) then shyps' | |
| 3061 | 536 | else (* FIXME tmp (since 1995/08/01) *) | 
| 2386 | 537 |               (warning ("Removed sort hypotheses: " ^
 | 
| 2962 | 538 | commas (map Sorts.str_of_sort (shyps' \\ sorts))); | 
| 2386 | 539 | warning "Let's hope these sorts are non-empty!"; | 
| 1238 | 540 | sorts)), | 
| 1529 | 541 | hyps = hyps, | 
| 542 | prop = prop} | |
| 1238 | 543 | end; | 
| 544 | ||
| 545 | ||
| 546 | (* implies_intr_shyps *) | |
| 547 | ||
| 548 | (*discharge all extra sort hypotheses*) | |
| 549 | fun implies_intr_shyps thm = | |
| 550 | (case extra_shyps thm of | |
| 551 | [] => thm | |
| 552 | | xshyps => | |
| 553 | let | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 554 |         val Thm {sign_ref, der, maxidx, shyps, hyps, prop} = thm;
 | 
| 2182 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 paulson parents: 
2177diff
changeset | 555 | val shyps' = ins_sort (logicS, shyps \\ xshyps); | 
| 1238 | 556 | val used_names = foldr add_term_tfree_names (prop :: hyps, []); | 
| 557 | val names = | |
| 558 | tl (variantlist (replicate (length xshyps + 1) "'", used_names)); | |
| 559 | val tfrees = map (TFree o rpair logicS) names; | |
| 560 | ||
| 561 | fun mk_insort (T, S) = map (Logic.mk_inclass o pair T) S; | |
| 2671 | 562 | val sort_hyps = List.concat (map2 mk_insort (tfrees, xshyps)); | 
| 1238 | 563 | in | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 564 |         Thm {sign_ref = sign_ref, 
 | 
| 2386 | 565 | der = infer_derivs (Implies_intr_shyps, [der]), | 
| 566 | maxidx = maxidx, | |
| 567 | shyps = shyps', | |
| 568 | hyps = hyps, | |
| 569 | prop = Logic.list_implies (sort_hyps, prop)} | |
| 1238 | 570 | end); | 
| 571 | ||
| 572 | ||
| 1529 | 573 | (** Axioms **) | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 574 | |
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 575 | (*look up the named axiom in the theory*) | 
| 3812 | 576 | fun get_axiom theory raw_name = | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 577 | let | 
| 4847 | 578 | val name = Sign.intern (Theory.sign_of theory) Theory.axiomK raw_name; | 
| 579 | ||
| 580 | fun get_ax [] = None | |
| 1529 | 581 | | get_ax (thy :: thys) = | 
| 4847 | 582 |           let val {sign, axioms, ...} = Theory.rep_theory thy in
 | 
| 583 | (case Symtab.lookup (axioms, name) of | |
| 584 | Some t => | |
| 585 | Some (fix_shyps [] [] | |
| 586 |                   (Thm {sign_ref = Sign.self_ref sign,
 | |
| 587 | der = infer_derivs (Axiom name, []), | |
| 588 | maxidx = maxidx_of_term t, | |
| 589 | shyps = [], | |
| 590 | hyps = [], | |
| 591 | prop = t})) | |
| 592 | | None => get_ax thys) | |
| 1529 | 593 | end; | 
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 594 | in | 
| 4847 | 595 | (case get_ax (theory :: Theory.ancestors_of theory) of | 
| 596 | Some thm => thm | |
| 597 |     | None => raise THEORY ("No axiom " ^ quote name, [theory]))
 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 598 | end; | 
| 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 599 | |
| 4847 | 600 | fun get_def thy name = get_axiom thy (name ^ "_def"); | 
| 601 | ||
| 1529 | 602 | |
| 776 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 603 | (*return additional axioms of this theory node*) | 
| 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 604 | fun axioms_of thy = | 
| 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 605 | map (fn (s, _) => (s, get_axiom thy s)) | 
| 3994 | 606 | (Symtab.dest (#axioms (rep_theory thy))); | 
| 776 
df8f91c0e57c
improved axioms_of: returns thms as the manual says;
 wenzelm parents: 
721diff
changeset | 607 | |
| 1597 
54ece585bf62
name_thm no longer takes a theory argument, as the
 paulson parents: 
1580diff
changeset | 608 | (*Attach a label to a theorem to make proof objects more readable*) | 
| 4018 | 609 | fun name_thm (name, th as Thm {sign_ref, der, maxidx, shyps, hyps, prop}) =
 | 
| 610 | (case der of | |
| 611 | Join (Theorem _, _) => th | |
| 612 | | Join (Axiom _, _) => th | |
| 613 |   | _ => Thm {sign_ref = sign_ref, der = Join (Theorem name, [der]),
 | |
| 614 | maxidx = maxidx, shyps = shyps, hyps = hyps, prop = prop}); | |
| 615 | ||
| 616 | fun name_of_thm (Thm {der, ...}) =
 | |
| 617 | (case der of | |
| 618 | Join (Theorem name, _) => name | |
| 4182 | 619 | | Join (Axiom name, _) => name | 
| 4018 | 620 | | _ => ""); | 
| 0 | 621 | |
| 622 | ||
| 1529 | 623 | (*Compression of theorems -- a separate rule, not integrated with the others, | 
| 624 | as it could be slow.*) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 625 | fun compress (Thm {sign_ref, der, maxidx, shyps, hyps, prop}) = 
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 626 |     Thm {sign_ref = sign_ref, 
 | 
| 2386 | 627 | der = der, (*No derivation recorded!*) | 
| 628 | maxidx = maxidx, | |
| 629 | shyps = shyps, | |
| 630 | hyps = map Term.compress_term hyps, | |
| 631 | prop = Term.compress_term prop}; | |
| 564 | 632 | |
| 387 
69f4356d915d
new datatype theory, supports 'draft theories' and incremental extension:
 wenzelm parents: 
309diff
changeset | 633 | |
| 2509 | 634 | |
| 1529 | 635 | (*** Meta rules ***) | 
| 0 | 636 | |
| 2147 | 637 | (*Check that term does not contain same var with different typing/sorting. | 
| 638 | If this check must be made, recalculate maxidx in hope of preventing its | |
| 639 | recurrence.*) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 640 | fun nodup_Vars (thm as Thm{sign_ref, der, maxidx, shyps, hyps, prop}) s =
 | 
| 2147 | 641 | (Sign.nodup_Vars prop; | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 642 |    Thm {sign_ref = sign_ref, 
 | 
| 2386 | 643 | der = der, | 
| 644 | maxidx = maxidx_of_term prop, | |
| 645 | shyps = shyps, | |
| 646 | hyps = hyps, | |
| 647 | prop = prop}) | |
| 2147 | 648 | handle TYPE(msg,Ts,ts) => raise TYPE(s^": "^msg,Ts,ts); | 
| 1495 
b8b54847c77f
Added check for duplicate vars with distinct types/sorts (nodup_Vars)
 nipkow parents: 
1493diff
changeset | 649 | |
| 1220 | 650 | (** 'primitive' rules **) | 
| 651 | ||
| 652 | (*discharge all assumptions t from ts*) | |
| 0 | 653 | val disch = gen_rem (op aconv); | 
| 654 | ||
| 1220 | 655 | (*The assumption rule A|-A in a theory*) | 
| 250 | 656 | fun assume ct : thm = | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 657 |   let val Cterm {sign_ref, t=prop, T, maxidx} = ct
 | 
| 250 | 658 | in if T<>propT then | 
| 659 |         raise THM("assume: assumptions must have type prop", 0, [])
 | |
| 0 | 660 | else if maxidx <> ~1 then | 
| 250 | 661 |         raise THM("assume: assumptions may not contain scheme variables",
 | 
| 662 | maxidx, []) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 663 |       else Thm{sign_ref   = sign_ref,
 | 
| 2386 | 664 | der = infer_derivs (Assume ct, []), | 
| 665 | maxidx = ~1, | |
| 666 | shyps = add_term_sorts(prop,[]), | |
| 667 | hyps = [prop], | |
| 668 | prop = prop} | |
| 0 | 669 | end; | 
| 670 | ||
| 1220 | 671 | (*Implication introduction | 
| 3529 | 672 | [A] | 
| 673 | : | |
| 674 | B | |
| 1220 | 675 | ------- | 
| 676 | A ==> B | |
| 677 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 678 | fun implies_intr cA (thB as Thm{sign_ref,der,maxidx,hyps,prop,...}) : thm =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 679 |   let val Cterm {sign_ref=sign_refA, t=A, T, maxidx=maxidxA} = cA
 | 
| 0 | 680 | in if T<>propT then | 
| 250 | 681 |         raise THM("implies_intr: assumptions must have type prop", 0, [thB])
 | 
| 1238 | 682 | else fix_shyps [thB] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 683 |         (Thm{sign_ref = Sign.merge_refs (sign_ref,sign_refA),  
 | 
| 2386 | 684 | der = infer_derivs (Implies_intr cA, [der]), | 
| 685 | maxidx = Int.max(maxidxA, maxidx), | |
| 686 | shyps = [], | |
| 687 | hyps = disch(hyps,A), | |
| 688 | prop = implies$A$prop}) | |
| 0 | 689 | handle TERM _ => | 
| 690 |         raise THM("implies_intr: incompatible signatures", 0, [thB])
 | |
| 691 | end; | |
| 692 | ||
| 1529 | 693 | |
| 1220 | 694 | (*Implication elimination | 
| 695 | A ==> B A | |
| 696 | ------------ | |
| 697 | B | |
| 698 | *) | |
| 0 | 699 | fun implies_elim thAB thA : thm = | 
| 1529 | 700 |     let val Thm{maxidx=maxA, der=derA, hyps=hypsA, prop=propA,...} = thA
 | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 701 |         and Thm{sign_ref, der, maxidx, hyps, prop,...} = thAB;
 | 
| 250 | 702 |         fun err(a) = raise THM("implies_elim: "^a, 0, [thAB,thA])
 | 
| 0 | 703 | in case prop of | 
| 250 | 704 | imp$A$B => | 
| 705 | if imp=implies andalso A aconv propA | |
| 1220 | 706 | then fix_shyps [thAB, thA] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 707 |                        (Thm{sign_ref= merge_thm_sgs(thAB,thA),
 | 
| 2386 | 708 | der = infer_derivs (Implies_elim, [der,derA]), | 
| 709 | maxidx = Int.max(maxA,maxidx), | |
| 710 | shyps = [], | |
| 711 | hyps = union_term(hypsA,hyps), (*dups suppressed*) | |
| 712 | prop = B}) | |
| 250 | 713 |                 else err("major premise")
 | 
| 714 |           | _ => err("major premise")
 | |
| 0 | 715 | end; | 
| 250 | 716 | |
| 1220 | 717 | (*Forall introduction. The Free or Var x must not be free in the hypotheses. | 
| 718 | A | |
| 719 | ----- | |
| 720 | !!x.A | |
| 721 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 722 | fun forall_intr cx (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 723 | let val x = term_of cx; | 
| 1238 | 724 | fun result(a,T) = fix_shyps [th] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 725 |         (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 726 | der = infer_derivs (Forall_intr cx, [der]), | 
| 727 | maxidx = maxidx, | |
| 728 | shyps = [], | |
| 729 | hyps = hyps, | |
| 730 | prop = all(T) $ Abs(a, T, abstract_over (x,prop))}) | |
| 0 | 731 | in case x of | 
| 250 | 732 | Free(a,T) => | 
| 733 | if exists (apl(x, Logic.occs)) hyps | |
| 734 |           then  raise THM("forall_intr: variable free in assumptions", 0, [th])
 | |
| 735 | else result(a,T) | |
| 0 | 736 | | Var((a,_),T) => result(a,T) | 
| 737 |       | _ => raise THM("forall_intr: not a variable", 0, [th])
 | |
| 738 | end; | |
| 739 | ||
| 1220 | 740 | (*Forall elimination | 
| 741 | !!x.A | |
| 742 | ------ | |
| 743 | A[t/x] | |
| 744 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 745 | fun forall_elim ct (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) : thm =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 746 |   let val Cterm {sign_ref=sign_reft, t, T, maxidx=maxt} = ct
 | 
| 0 | 747 | in case prop of | 
| 2386 | 748 |         Const("all",Type("fun",[Type("fun",[qary,_]),_])) $ A =>
 | 
| 749 | if T<>qary then | |
| 750 |               raise THM("forall_elim: type mismatch", 0, [th])
 | |
| 751 | else let val thm = fix_shyps [th] [] | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 752 |                     (Thm{sign_ref= Sign.merge_refs(sign_ref,sign_reft),
 | 
| 2386 | 753 | der = infer_derivs (Forall_elim ct, [der]), | 
| 754 | maxidx = Int.max(maxidx, maxt), | |
| 755 | shyps = [], | |
| 756 | hyps = hyps, | |
| 757 | prop = betapply(A,t)}) | |
| 758 | in if maxt >= 0 andalso maxidx >= 0 | |
| 759 | then nodup_Vars thm "forall_elim" | |
| 760 | else thm (*no new Vars: no expensive check!*) | |
| 761 | end | |
| 2147 | 762 |       | _ => raise THM("forall_elim: not quantified", 0, [th])
 | 
| 0 | 763 | end | 
| 764 | handle TERM _ => | |
| 250 | 765 |          raise THM("forall_elim: incompatible signatures", 0, [th]);
 | 
| 0 | 766 | |
| 767 | ||
| 1220 | 768 | (* Equality *) | 
| 0 | 769 | |
| 770 | (*The reflexivity rule: maps t to the theorem t==t *) | |
| 250 | 771 | fun reflexive ct = | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 772 |   let val Cterm {sign_ref, t, T, maxidx} = ct
 | 
| 1238 | 773 | in fix_shyps [] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 774 |        (Thm{sign_ref= sign_ref, 
 | 
| 2386 | 775 | der = infer_derivs (Reflexive ct, []), | 
| 776 | shyps = [], | |
| 777 | hyps = [], | |
| 778 | maxidx = maxidx, | |
| 779 | prop = Logic.mk_equals(t,t)}) | |
| 0 | 780 | end; | 
| 781 | ||
| 782 | (*The symmetry rule | |
| 1220 | 783 | t==u | 
| 784 | ---- | |
| 785 | u==t | |
| 786 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 787 | fun symmetric (th as Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 0 | 788 | case prop of | 
| 789 |       (eq as Const("==",_)) $ t $ u =>
 | |
| 1238 | 790 | (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 791 |           Thm{sign_ref = sign_ref,
 | 
| 2386 | 792 | der = infer_derivs (Symmetric, [der]), | 
| 793 | maxidx = maxidx, | |
| 794 | shyps = shyps, | |
| 795 | hyps = hyps, | |
| 796 | prop = eq$u$t} | |
| 0 | 797 |     | _ => raise THM("symmetric", 0, [th]);
 | 
| 798 | ||
| 799 | (*The transitive rule | |
| 1220 | 800 | t1==u u==t2 | 
| 801 | -------------- | |
| 802 | t1==t2 | |
| 803 | *) | |
| 0 | 804 | fun transitive th1 th2 = | 
| 1529 | 805 |   let val Thm{der=der1, maxidx=max1, hyps=hyps1, prop=prop1,...} = th1
 | 
| 806 |       and Thm{der=der2, maxidx=max2, hyps=hyps2, prop=prop2,...} = th2;
 | |
| 0 | 807 |       fun err(msg) = raise THM("transitive: "^msg, 0, [th1,th2])
 | 
| 808 | in case (prop1,prop2) of | |
| 809 |        ((eq as Const("==",_)) $ t1 $ u, Const("==",_) $ u' $ t2) =>
 | |
| 1634 | 810 | if not (u aconv u') then err"middle term" | 
| 811 | else let val thm = | |
| 1220 | 812 | fix_shyps [th1, th2] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 813 |                 (Thm{sign_ref= merge_thm_sgs(th1,th2), 
 | 
| 2386 | 814 | der = infer_derivs (Transitive, [der1, der2]), | 
| 2147 | 815 | maxidx = Int.max(max1,max2), | 
| 2386 | 816 | shyps = [], | 
| 817 | hyps = union_term(hyps1,hyps2), | |
| 818 | prop = eq$t1$t2}) | |
| 2139 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 paulson parents: 
2047diff
changeset | 819 | in if max1 >= 0 andalso max2 >= 0 | 
| 2147 | 820 | then nodup_Vars thm "transitive" | 
| 821 | else thm (*no new Vars: no expensive check!*) | |
| 2139 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 paulson parents: 
2047diff
changeset | 822 | end | 
| 0 | 823 | | _ => err"premises" | 
| 824 | end; | |
| 825 | ||
| 1160 | 826 | (*Beta-conversion: maps (%x.t)(u) to the theorem (%x.t)(u) == t[u/x] *) | 
| 250 | 827 | fun beta_conversion ct = | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 828 |   let val Cterm {sign_ref, t, T, maxidx} = ct
 | 
| 0 | 829 | in case t of | 
| 1238 | 830 | Abs(_,_,bodt) $ u => fix_shyps [] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 831 |             (Thm{sign_ref = sign_ref,  
 | 
| 2386 | 832 | der = infer_derivs (Beta_conversion ct, []), | 
| 833 | maxidx = maxidx, | |
| 834 | shyps = [], | |
| 835 | hyps = [], | |
| 836 | prop = Logic.mk_equals(t, subst_bound (u,bodt))}) | |
| 250 | 837 |         | _ =>  raise THM("beta_conversion: not a redex", 0, [])
 | 
| 0 | 838 | end; | 
| 839 | ||
| 840 | (*The extensionality rule (proviso: x not free in f, g, or hypotheses) | |
| 1220 | 841 | f(x) == g(x) | 
| 842 | ------------ | |
| 843 | f == g | |
| 844 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 845 | fun extensional (th as Thm{sign_ref, der, maxidx,shyps,hyps,prop}) =
 | 
| 0 | 846 | case prop of | 
| 847 |     (Const("==",_)) $ (f$x) $ (g$y) =>
 | |
| 250 | 848 |       let fun err(msg) = raise THM("extensional: "^msg, 0, [th])
 | 
| 0 | 849 | in (if x<>y then err"different variables" else | 
| 850 | case y of | |
| 250 | 851 | Free _ => | 
| 852 | if exists (apl(y, Logic.occs)) (f::g::hyps) | |
| 853 | then err"variable free in hyps or functions" else () | |
| 854 | | Var _ => | |
| 855 | if Logic.occs(y,f) orelse Logic.occs(y,g) | |
| 856 | then err"variable free in functions" else () | |
| 857 | | _ => err"not a variable"); | |
| 1238 | 858 | (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 859 |           Thm{sign_ref = sign_ref,
 | 
| 2386 | 860 | der = infer_derivs (Extensional, [der]), | 
| 861 | maxidx = maxidx, | |
| 862 | shyps = shyps, | |
| 863 | hyps = hyps, | |
| 1529 | 864 | prop = Logic.mk_equals(f,g)} | 
| 0 | 865 | end | 
| 866 |  | _ =>  raise THM("extensional: premise", 0, [th]);
 | |
| 867 | ||
| 868 | (*The abstraction rule. The Free or Var x must not be free in the hypotheses. | |
| 869 | The bound variable will be named "a" (since x will be something like x320) | |
| 1220 | 870 | t == u | 
| 871 | ------------ | |
| 872 | %x.t == %x.u | |
| 873 | *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 874 | fun abstract_rule a cx (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 229 
4002c4cd450c
Pure:  MAJOR CHANGE.  Moved ML types ctyp and cterm and their associated
 lcp parents: 
225diff
changeset | 875 | let val x = term_of cx; | 
| 250 | 876 | val (t,u) = Logic.dest_equals prop | 
| 877 | handle TERM _ => | |
| 878 |                 raise THM("abstract_rule: premise not an equality", 0, [th])
 | |
| 1238 | 879 | fun result T = fix_shyps [th] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 880 |           (Thm{sign_ref = sign_ref,
 | 
| 2386 | 881 | der = infer_derivs (Abstract_rule (a,cx), [der]), | 
| 882 | maxidx = maxidx, | |
| 883 | shyps = [], | |
| 884 | hyps = hyps, | |
| 885 | prop = Logic.mk_equals(Abs(a, T, abstract_over (x,t)), | |
| 886 | Abs(a, T, abstract_over (x,u)))}) | |
| 0 | 887 | in case x of | 
| 250 | 888 | Free(_,T) => | 
| 889 | if exists (apl(x, Logic.occs)) hyps | |
| 890 |          then raise THM("abstract_rule: variable free in assumptions", 0, [th])
 | |
| 891 | else result T | |
| 0 | 892 | | Var(_,T) => result T | 
| 893 |       | _ => raise THM("abstract_rule: not a variable", 0, [th])
 | |
| 894 | end; | |
| 895 | ||
| 896 | (*The combination rule | |
| 3529 | 897 | f == g t == u | 
| 898 | -------------- | |
| 899 | f(t) == g(u) | |
| 1220 | 900 | *) | 
| 0 | 901 | fun combination th1 th2 = | 
| 1529 | 902 |   let val Thm{der=der1, maxidx=max1, shyps=shyps1, hyps=hyps1, 
 | 
| 2386 | 903 | prop=prop1,...} = th1 | 
| 1529 | 904 |       and Thm{der=der2, maxidx=max2, shyps=shyps2, hyps=hyps2, 
 | 
| 2386 | 905 | prop=prop2,...} = th2 | 
| 1836 
861e29c7cada
Added type-checking to rule "combination".  This corrects a fault
 paulson parents: 
1802diff
changeset | 906 | fun chktypes (f,t) = | 
| 2386 | 907 | (case fastype_of f of | 
| 908 |                 Type("fun",[T1,T2]) => 
 | |
| 909 | if T1 <> fastype_of t then | |
| 910 |                          raise THM("combination: types", 0, [th1,th2])
 | |
| 911 | else () | |
| 912 |                 | _ => raise THM("combination: not function type", 0, 
 | |
| 913 | [th1,th2])) | |
| 1495 
b8b54847c77f
Added check for duplicate vars with distinct types/sorts (nodup_Vars)
 nipkow parents: 
1493diff
changeset | 914 | in case (prop1,prop2) of | 
| 0 | 915 |        (Const("==",_) $ f $ g, Const("==",_) $ t $ u) =>
 | 
| 1836 
861e29c7cada
Added type-checking to rule "combination".  This corrects a fault
 paulson parents: 
1802diff
changeset | 916 | let val _ = chktypes (f,t) | 
| 2386 | 917 | val thm = (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 918 |                         Thm{sign_ref = merge_thm_sgs(th1,th2), 
 | 
| 2386 | 919 | der = infer_derivs (Combination, [der1, der2]), | 
| 920 | maxidx = Int.max(max1,max2), | |
| 921 | shyps = union_sort(shyps1,shyps2), | |
| 922 | hyps = union_term(hyps1,hyps2), | |
| 923 | prop = Logic.mk_equals(f$t, g$u)} | |
| 2139 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 paulson parents: 
2047diff
changeset | 924 | in if max1 >= 0 andalso max2 >= 0 | 
| 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 paulson parents: 
2047diff
changeset | 925 | then nodup_Vars thm "combination" | 
| 2386 | 926 | else thm (*no new Vars: no expensive check!*) | 
| 2139 
2c59b204b540
Only calls nodup_Vars if really necessary.  We get a speedup of nearly 6%
 paulson parents: 
2047diff
changeset | 927 | end | 
| 0 | 928 |      | _ =>  raise THM("combination: premises", 0, [th1,th2])
 | 
| 929 | end; | |
| 930 | ||
| 931 | ||
| 932 | (* Equality introduction | |
| 3529 | 933 | A ==> B B ==> A | 
| 934 | ---------------- | |
| 935 | A == B | |
| 1220 | 936 | *) | 
| 0 | 937 | fun equal_intr th1 th2 = | 
| 1529 | 938 |   let val Thm{der=der1,maxidx=max1, shyps=shyps1, hyps=hyps1, 
 | 
| 2386 | 939 | prop=prop1,...} = th1 | 
| 1529 | 940 |       and Thm{der=der2, maxidx=max2, shyps=shyps2, hyps=hyps2, 
 | 
| 2386 | 941 | prop=prop2,...} = th2; | 
| 1529 | 942 |       fun err(msg) = raise THM("equal_intr: "^msg, 0, [th1,th2])
 | 
| 943 | in case (prop1,prop2) of | |
| 944 |        (Const("==>",_) $ A $ B, Const("==>",_) $ B' $ A')  =>
 | |
| 2386 | 945 | if A aconv A' andalso B aconv B' | 
| 946 | then | |
| 947 | (*no fix_shyps*) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 948 |               Thm{sign_ref = merge_thm_sgs(th1,th2),
 | 
| 2386 | 949 | der = infer_derivs (Equal_intr, [der1, der2]), | 
| 950 | maxidx = Int.max(max1,max2), | |
| 951 | shyps = union_sort(shyps1,shyps2), | |
| 952 | hyps = union_term(hyps1,hyps2), | |
| 953 | prop = Logic.mk_equals(A,B)} | |
| 954 | else err"not equal" | |
| 1529 | 955 | | _ => err"premises" | 
| 956 | end; | |
| 957 | ||
| 958 | ||
| 959 | (*The equal propositions rule | |
| 3529 | 960 | A == B A | 
| 1529 | 961 | --------- | 
| 962 | B | |
| 963 | *) | |
| 964 | fun equal_elim th1 th2 = | |
| 965 |   let val Thm{der=der1, maxidx=max1, hyps=hyps1, prop=prop1,...} = th1
 | |
| 966 |       and Thm{der=der2, maxidx=max2, hyps=hyps2, prop=prop2,...} = th2;
 | |
| 967 |       fun err(msg) = raise THM("equal_elim: "^msg, 0, [th1,th2])
 | |
| 968 | in case prop1 of | |
| 969 |        Const("==",_) $ A $ B =>
 | |
| 970 | if not (prop2 aconv A) then err"not equal" else | |
| 971 | fix_shyps [th1, th2] [] | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 972 |               (Thm{sign_ref= merge_thm_sgs(th1,th2), 
 | 
| 2386 | 973 | der = infer_derivs (Equal_elim, [der1, der2]), | 
| 974 | maxidx = Int.max(max1,max2), | |
| 975 | shyps = [], | |
| 976 | hyps = union_term(hyps1,hyps2), | |
| 977 | prop = B}) | |
| 1529 | 978 | | _ => err"major premise" | 
| 979 | end; | |
| 0 | 980 | |
| 1220 | 981 | |
| 982 | ||
| 0 | 983 | (**** Derived rules ****) | 
| 984 | ||
| 1503 | 985 | (*Discharge all hypotheses. Need not verify cterms or call fix_shyps. | 
| 0 | 986 | Repeated hypotheses are discharged only once; fold cannot do this*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 987 | fun implies_intr_hyps (Thm{sign_ref, der, maxidx, shyps, hyps=A::As, prop}) =
 | 
| 1238 | 988 | implies_intr_hyps (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 989 |             (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 990 | der = infer_derivs (Implies_intr_hyps, [der]), | 
| 991 | maxidx = maxidx, | |
| 992 | shyps = shyps, | |
| 1529 | 993 | hyps = disch(As,A), | 
| 2386 | 994 | prop = implies$A$prop}) | 
| 0 | 995 | | implies_intr_hyps th = th; | 
| 996 | ||
| 997 | (*Smash" unifies the list of term pairs leaving no flex-flex pairs. | |
| 250 | 998 | Instantiates the theorem and deletes trivial tpairs. | 
| 0 | 999 | Resulting sequence may contain multiple elements if the tpairs are | 
| 1000 | not all flex-flex. *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1001 | fun flexflex_rule (th as Thm{sign_ref, der, maxidx, hyps, prop,...}) =
 | 
| 250 | 1002 | let fun newthm env = | 
| 1529 | 1003 | if Envir.is_empty env then th | 
| 1004 | else | |
| 250 | 1005 | let val (tpairs,horn) = | 
| 1006 | Logic.strip_flexpairs (Envir.norm_term env prop) | |
| 1007 | (*Remove trivial tpairs, of the form t=t*) | |
| 1008 | val distpairs = filter (not o op aconv) tpairs | |
| 1009 | val newprop = Logic.list_flexpairs(distpairs, horn) | |
| 1220 | 1010 | in fix_shyps [th] (env_codT env) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1011 |                 (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1012 | der = infer_derivs (Flexflex_rule env, [der]), | 
| 1013 | maxidx = maxidx_of_term newprop, | |
| 1014 | shyps = [], | |
| 1015 | hyps = hyps, | |
| 1016 | prop = newprop}) | |
| 250 | 1017 | end; | 
| 0 | 1018 | val (tpairs,_) = Logic.strip_flexpairs prop | 
| 4270 | 1019 | in Seq.map newthm | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1020 | (Unify.smash_unifiers(Sign.deref sign_ref, Envir.empty maxidx, tpairs)) | 
| 0 | 1021 | end; | 
| 1022 | ||
| 1023 | (*Instantiation of Vars | |
| 1220 | 1024 | A | 
| 1025 | ------------------- | |
| 1026 | A[t1/v1,....,tn/vn] | |
| 1027 | *) | |
| 0 | 1028 | |
| 1029 | (*Check that all the terms are Vars and are distinct*) | |
| 1030 | fun instl_ok ts = forall is_Var ts andalso null(findrep ts); | |
| 1031 | ||
| 1032 | (*For instantiate: process pair of cterms, merge theories*) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1033 | fun add_ctpair ((ct,cu), (sign_ref,tpairs)) = | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1034 |   let val Cterm {sign_ref=sign_reft, t=t, T= T, ...} = ct
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1035 |       and Cterm {sign_ref=sign_refu, t=u, T= U, ...} = cu
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1036 | in | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1037 | if T=U then | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1038 | (Sign.merge_refs (sign_ref, Sign.merge_refs (sign_reft, sign_refu)), (t,u)::tpairs) | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1039 |     else raise TYPE("add_ctpair", [T,U], [t,u])
 | 
| 0 | 1040 | end; | 
| 1041 | ||
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1042 | fun add_ctyp ((v,ctyp), (sign_ref',vTs)) = | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1043 |   let val Ctyp {T,sign_ref} = ctyp
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1044 | in (Sign.merge_refs(sign_ref,sign_ref'), (v,T)::vTs) end; | 
| 0 | 1045 | |
| 1046 | (*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. | |
| 1047 | Instantiates distinct Vars by terms of same type. | |
| 1048 | Normalizes the new theorem! *) | |
| 1529 | 1049 | fun instantiate ([], []) th = th | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1050 |   | instantiate (vcTs,ctpairs)  (th as Thm{sign_ref,der,maxidx,hyps,prop,...}) =
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1051 | let val (newsign_ref,tpairs) = foldr add_ctpair (ctpairs, (sign_ref,[])); | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1052 | val (newsign_ref,vTs) = foldr add_ctyp (vcTs, (newsign_ref,[])); | 
| 250 | 1053 | val newprop = | 
| 1054 | Envir.norm_term (Envir.empty 0) | |
| 1055 | (subst_atomic tpairs | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1056 | (Type.inst_term_tvars(Sign.tsig_of (Sign.deref newsign_ref),vTs) prop)) | 
| 1220 | 1057 | val newth = | 
| 1058 | fix_shyps [th] (map snd vTs) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1059 |               (Thm{sign_ref = newsign_ref, 
 | 
| 2386 | 1060 | der = infer_derivs (Instantiate(vcTs,ctpairs), [der]), | 
| 1061 | maxidx = maxidx_of_term newprop, | |
| 1062 | shyps = [], | |
| 1063 | hyps = hyps, | |
| 1064 | prop = newprop}) | |
| 250 | 1065 | in if not(instl_ok(map #1 tpairs)) | 
| 193 | 1066 |       then raise THM("instantiate: variables not distinct", 0, [th])
 | 
| 1067 | else if not(null(findrep(map #1 vTs))) | |
| 1068 |       then raise THM("instantiate: type variables not distinct", 0, [th])
 | |
| 2147 | 1069 | else nodup_Vars newth "instantiate" | 
| 0 | 1070 | end | 
| 250 | 1071 | handle TERM _ => | 
| 0 | 1072 |            raise THM("instantiate: incompatible signatures",0,[th])
 | 
| 2671 | 1073 |        | TYPE (msg,_,_) => raise THM("instantiate: type conflict: " ^ msg, 
 | 
| 1074 | 0, [th]); | |
| 0 | 1075 | |
| 1076 | (*The trivial implication A==>A, justified by assume and forall rules. | |
| 1077 | A can contain Vars, not so for assume! *) | |
| 250 | 1078 | fun trivial ct : thm = | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1079 |   let val Cterm {sign_ref, t=A, T, maxidx} = ct
 | 
| 250 | 1080 | in if T<>propT then | 
| 1081 |             raise THM("trivial: the term must have type prop", 0, [])
 | |
| 1238 | 1082 | else fix_shyps [] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1083 |         (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1084 | der = infer_derivs (Trivial ct, []), | 
| 1085 | maxidx = maxidx, | |
| 1086 | shyps = [], | |
| 1087 | hyps = [], | |
| 1088 | prop = implies$A$A}) | |
| 0 | 1089 | end; | 
| 1090 | ||
| 1503 | 1091 | (*Axiom-scheme reflecting signature contents: "OFCLASS(?'a::c, c_class)" *) | 
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1092 | fun class_triv thy c = | 
| 1529 | 1093 | let val sign = sign_of thy; | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1094 |       val Cterm {sign_ref, t, maxidx, ...} =
 | 
| 2386 | 1095 |           cterm_of sign (Logic.mk_inclass (TVar (("'a", 0), [c]), c))
 | 
| 1096 |             handle TERM (msg, _) => raise THM ("class_triv: " ^ msg, 0, []);
 | |
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1097 | in | 
| 1238 | 1098 | fix_shyps [] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1099 |       (Thm {sign_ref = sign_ref, 
 | 
| 4182 | 1100 | der = infer_derivs (Class_triv c, []), | 
| 2386 | 1101 | maxidx = maxidx, | 
| 1102 | shyps = [], | |
| 1103 | hyps = [], | |
| 1104 | prop = t}) | |
| 399 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1105 | end; | 
| 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1106 | |
| 
86cc2b98f9e0
added class_triv: theory -> class -> thm (for axclasses);
 wenzelm parents: 
387diff
changeset | 1107 | |
| 0 | 1108 | (* Replace all TFrees not in the hyps by new TVars *) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1109 | fun varifyT(Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 0 | 1110 | let val tfrees = foldr add_term_tfree_names (hyps,[]) | 
| 1634 | 1111 | in let val thm = (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1112 |     Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1113 | der = infer_derivs (VarifyT, [der]), | 
| 1114 | maxidx = Int.max(0,maxidx), | |
| 1115 | shyps = shyps, | |
| 1116 | hyps = hyps, | |
| 1529 | 1117 | prop = Type.varify(prop,tfrees)} | 
| 2147 | 1118 | in nodup_Vars thm "varifyT" end | 
| 1634 | 1119 | (* this nodup_Vars check can be removed if thms are guaranteed not to contain | 
| 1120 | duplicate TVars with differnt sorts *) | |
| 0 | 1121 | end; | 
| 1122 | ||
| 1123 | (* Replace all TVars by new TFrees *) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1124 | fun freezeT(Thm{sign_ref,der,maxidx,shyps,hyps,prop}) =
 | 
| 3410 | 1125 | let val (prop',_) = Type.freeze_thaw prop | 
| 1238 | 1126 | in (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1127 |     Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1128 | der = infer_derivs (FreezeT, [der]), | 
| 1129 | maxidx = maxidx_of_term prop', | |
| 1130 | shyps = shyps, | |
| 1131 | hyps = hyps, | |
| 1529 | 1132 | prop = prop'} | 
| 1220 | 1133 | end; | 
| 0 | 1134 | |
| 1135 | ||
| 1136 | (*** Inference rules for tactics ***) | |
| 1137 | ||
| 1138 | (*Destruct proof state into constraints, other goals, goal(i), rest *) | |
| 1139 | fun dest_state (state as Thm{prop,...}, i) =
 | |
| 1140 | let val (tpairs,horn) = Logic.strip_flexpairs prop | |
| 1141 | in case Logic.strip_prems(i, [], horn) of | |
| 1142 | (B::rBs, C) => (tpairs, rev rBs, B, C) | |
| 1143 |         | _ => raise THM("dest_state", i, [state])
 | |
| 1144 | end | |
| 1145 |   handle TERM _ => raise THM("dest_state", i, [state]);
 | |
| 1146 | ||
| 309 | 1147 | (*Increment variables and parameters of orule as required for | 
| 0 | 1148 | resolution with goal i of state. *) | 
| 1149 | fun lift_rule (state, i) orule = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1150 |   let val Thm{shyps=sshyps, prop=sprop, maxidx=smax, sign_ref=ssign_ref,...} = state
 | 
| 0 | 1151 | val (Bi::_, _) = Logic.strip_prems(i, [], Logic.skip_flexpairs sprop) | 
| 1529 | 1152 |         handle TERM _ => raise THM("lift_rule", i, [orule,state])
 | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1153 |       val ct_Bi = Cterm {sign_ref=ssign_ref, maxidx=smax, T=propT, t=Bi}
 | 
| 1529 | 1154 | val (lift_abs,lift_all) = Logic.lift_fns(Bi,smax+1) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1155 |       val (Thm{sign_ref, der, maxidx,shyps,hyps,prop}) = orule
 | 
| 0 | 1156 | val (tpairs,As,B) = Logic.strip_horn prop | 
| 1238 | 1157 | in (*no fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1158 |       Thm{sign_ref = merge_thm_sgs(state,orule),
 | 
| 2386 | 1159 | der = infer_derivs (Lift_rule(ct_Bi, i), [der]), | 
| 1160 | maxidx = maxidx+smax+1, | |
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 1161 | shyps=union_sort(sshyps,shyps), | 
| 2386 | 1162 | hyps=hyps, | 
| 1529 | 1163 | prop = Logic.rule_of (map (pairself lift_abs) tpairs, | 
| 2386 | 1164 | map lift_all As, | 
| 1165 | lift_all B)} | |
| 0 | 1166 | end; | 
| 1167 | ||
| 1168 | (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) | |
| 1169 | fun assumption i state = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1170 |   let val Thm{sign_ref,der,maxidx,hyps,prop,...} = state;
 | 
| 0 | 1171 | val (tpairs, Bs, Bi, C) = dest_state(state,i) | 
| 250 | 1172 |       fun newth (env as Envir.Envir{maxidx, ...}, tpairs) =
 | 
| 1220 | 1173 | fix_shyps [state] (env_codT env) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1174 |           (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1175 | der = infer_derivs (Assumption (i, Some env), [der]), | 
| 1176 | maxidx = maxidx, | |
| 1177 | shyps = [], | |
| 1178 | hyps = hyps, | |
| 1179 | prop = | |
| 1180 | if Envir.is_empty env then (*avoid wasted normalizations*) | |
| 1181 | Logic.rule_of (tpairs, Bs, C) | |
| 1182 | else (*normalize the new rule fully*) | |
| 1183 | Envir.norm_term env (Logic.rule_of (tpairs, Bs, C))}); | |
| 4270 | 1184 | fun addprfs [] = Seq.empty | 
| 1185 | | addprfs ((t,u)::apairs) = Seq.make (fn()=> Seq.pull | |
| 1186 | (Seq.mapp newth | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1187 | (Unify.unifiers(Sign.deref sign_ref,Envir.empty maxidx, (t,u)::tpairs)) | 
| 250 | 1188 | (addprfs apairs))) | 
| 0 | 1189 | in addprfs (Logic.assum_pairs Bi) end; | 
| 1190 | ||
| 250 | 1191 | (*Solve subgoal Bi of proof state B1...Bn/C by assumption. | 
| 0 | 1192 | Checks if Bi's conclusion is alpha-convertible to one of its assumptions*) | 
| 1193 | fun eq_assumption i state = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1194 |   let val Thm{sign_ref,der,maxidx,hyps,prop,...} = state;
 | 
| 0 | 1195 | val (tpairs, Bs, Bi, C) = dest_state(state,i) | 
| 1196 | in if exists (op aconv) (Logic.assum_pairs Bi) | |
| 1220 | 1197 | then fix_shyps [state] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1198 |              (Thm{sign_ref = sign_ref, 
 | 
| 2386 | 1199 | der = infer_derivs (Assumption (i,None), [der]), | 
| 1200 | maxidx = maxidx, | |
| 1201 | shyps = [], | |
| 1202 | hyps = hyps, | |
| 1203 | prop = Logic.rule_of(tpairs, Bs, C)}) | |
| 0 | 1204 |       else  raise THM("eq_assumption", 0, [state])
 | 
| 1205 | end; | |
| 1206 | ||
| 1207 | ||
| 2671 | 1208 | (*For rotate_tac: fast rotation of assumptions of subgoal i*) | 
| 1209 | fun rotate_rule k i state = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1210 |   let val Thm{sign_ref,der,maxidx,hyps,prop,shyps} = state;
 | 
| 2671 | 1211 | val (tpairs, Bs, Bi, C) = dest_state(state,i) | 
| 1212 | val params = Logic.strip_params Bi | |
| 1213 | and asms = Logic.strip_assums_hyp Bi | |
| 1214 | and concl = Logic.strip_assums_concl Bi | |
| 1215 | val n = length asms | |
| 1216 | fun rot m = if 0=m orelse m=n then Bi | |
| 1217 | else if 0<m andalso m<n | |
| 1218 | then list_all | |
| 1219 | (params, | |
| 1220 | Logic.list_implies(List.drop(asms, m) @ | |
| 1221 | List.take(asms, m), | |
| 1222 | concl)) | |
| 1223 | 		   else raise THM("rotate_rule", m, [state])
 | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1224 |   in  Thm{sign_ref = sign_ref, 
 | 
| 2671 | 1225 | der = infer_derivs (Rotate_rule (k,i), [der]), | 
| 1226 | maxidx = maxidx, | |
| 1227 | shyps = shyps, | |
| 1228 | hyps = hyps, | |
| 1229 | prop = Logic.rule_of(tpairs, Bs@[rot (if k<0 then n+k else k)], C)} | |
| 1230 | end; | |
| 1231 | ||
| 1232 | ||
| 0 | 1233 | (** User renaming of parameters in a subgoal **) | 
| 1234 | ||
| 1235 | (*Calls error rather than raising an exception because it is intended | |
| 1236 | for top-level use -- exception handling would not make sense here. | |
| 1237 | The names in cs, if distinct, are used for the innermost parameters; | |
| 1238 | preceding parameters may be renamed to make all params distinct.*) | |
| 1239 | fun rename_params_rule (cs, i) state = | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1240 |   let val Thm{sign_ref,der,maxidx,hyps,...} = state
 | 
| 0 | 1241 | val (tpairs, Bs, Bi, C) = dest_state(state,i) | 
| 1242 | val iparams = map #1 (Logic.strip_params Bi) | |
| 1243 | val short = length iparams - length cs | |
| 250 | 1244 | val newnames = | 
| 1245 | if short<0 then error"More names than abstractions!" | |
| 1246 | else variantlist(take (short,iparams), cs) @ cs | |
| 3037 
99ed078e6ae7
rename_params_rule used to check if the new name clashed with a free name in
 nipkow parents: 
3012diff
changeset | 1247 | val freenames = map (#1 o dest_Free) (term_frees Bi) | 
| 0 | 1248 | val newBi = Logic.list_rename_params (newnames, Bi) | 
| 250 | 1249 | in | 
| 0 | 1250 | case findrep cs of | 
| 3565 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 paulson parents: 
3558diff
changeset | 1251 |      c::_ => (warning ("Can't rename.  Bound variables not distinct: " ^ c); 
 | 
| 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 paulson parents: 
3558diff
changeset | 1252 | state) | 
| 1576 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 berghofe parents: 
1569diff
changeset | 1253 | | [] => (case cs inter_string freenames of | 
| 3565 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 paulson parents: 
3558diff
changeset | 1254 |        a::_ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); 
 | 
| 
c64410e701fb
Now rename_params_rule merely issues warnings--and does nothing--if the
 paulson parents: 
3558diff
changeset | 1255 | state) | 
| 1220 | 1256 | | [] => fix_shyps [state] [] | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1257 |                 (Thm{sign_ref = sign_ref,
 | 
| 2386 | 1258 | der = infer_derivs (Rename_params_rule(cs,i), [der]), | 
| 1259 | maxidx = maxidx, | |
| 1260 | shyps = [], | |
| 1261 | hyps = hyps, | |
| 1262 | prop = Logic.rule_of(tpairs, Bs@[newBi], C)})) | |
| 0 | 1263 | end; | 
| 1264 | ||
| 1265 | (*** Preservation of bound variable names ***) | |
| 1266 | ||
| 250 | 1267 | (*Scan a pair of terms; while they are similar, | 
| 0 | 1268 | accumulate corresponding bound vars in "al"*) | 
| 1238 | 1269 | fun match_bvs(Abs(x,_,s),Abs(y,_,t), al) = | 
| 1195 
686e3eb613b9
match_bvs no longer puts a name in the alist if it is null ("")
 lcp parents: 
1160diff
changeset | 1270 | match_bvs(s, t, if x="" orelse y="" then al | 
| 1238 | 1271 | else (x,y)::al) | 
| 0 | 1272 | | match_bvs(f$s, g$t, al) = match_bvs(f,g,match_bvs(s,t,al)) | 
| 1273 | | match_bvs(_,_,al) = al; | |
| 1274 | ||
| 1275 | (* strip abstractions created by parameters *) | |
| 1276 | fun match_bvars((s,t),al) = match_bvs(strip_abs_body s, strip_abs_body t, al); | |
| 1277 | ||
| 1278 | ||
| 250 | 1279 | (* strip_apply f A(,B) strips off all assumptions/parameters from A | 
| 0 | 1280 | introduced by lifting over B, and applies f to remaining part of A*) | 
| 1281 | fun strip_apply f = | |
| 1282 |   let fun strip(Const("==>",_)$ A1 $ B1,
 | |
| 250 | 1283 |                 Const("==>",_)$ _  $ B2) = implies $ A1 $ strip(B1,B2)
 | 
| 1284 |         | strip((c as Const("all",_)) $ Abs(a,T,t1),
 | |
| 1285 |                       Const("all",_)  $ Abs(_,_,t2)) = c$Abs(a,T,strip(t1,t2))
 | |
| 1286 | | strip(A,_) = f A | |
| 0 | 1287 | in strip end; | 
| 1288 | ||
| 1289 | (*Use the alist to rename all bound variables and some unknowns in a term | |
| 1290 | dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); | |
| 1291 | Preserves unknowns in tpairs and on lhs of dpairs. *) | |
| 1292 | fun rename_bvs([],_,_,_) = I | |
| 1293 | | rename_bvs(al,dpairs,tpairs,B) = | |
| 250 | 1294 | let val vars = foldr add_term_vars | 
| 1295 | (map fst dpairs @ map fst tpairs @ map snd tpairs, []) | |
| 1296 | (*unknowns appearing elsewhere be preserved!*) | |
| 1297 | val vids = map (#1 o #1 o dest_Var) vars; | |
| 1298 | fun rename(t as Var((x,i),T)) = | |
| 1299 | (case assoc(al,x) of | |
| 1576 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 berghofe parents: 
1569diff
changeset | 1300 | Some(y) => if x mem_string vids orelse y mem_string vids then t | 
| 250 | 1301 | else Var((y,i),T) | 
| 1302 | | None=> t) | |
| 0 | 1303 | | rename(Abs(x,T,t)) = | 
| 1576 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 berghofe parents: 
1569diff
changeset | 1304 | Abs(case assoc_string(al,x) of Some(y) => y | None => x, | 
| 250 | 1305 | T, rename t) | 
| 0 | 1306 | | rename(f$t) = rename f $ rename t | 
| 1307 | | rename(t) = t; | |
| 250 | 1308 | fun strip_ren Ai = strip_apply rename (Ai,B) | 
| 0 | 1309 | in strip_ren end; | 
| 1310 | ||
| 1311 | (*Function to rename bounds/unknowns in the argument, lifted over B*) | |
| 1312 | fun rename_bvars(dpairs, tpairs, B) = | |
| 250 | 1313 | rename_bvs(foldr match_bvars (dpairs,[]), dpairs, tpairs, B); | 
| 0 | 1314 | |
| 1315 | ||
| 1316 | (*** RESOLUTION ***) | |
| 1317 | ||
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1318 | (** Lifting optimizations **) | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1319 | |
| 0 | 1320 | (*strip off pairs of assumptions/parameters in parallel -- they are | 
| 1321 | identical because of lifting*) | |
| 250 | 1322 | fun strip_assums2 (Const("==>", _) $ _ $ B1,
 | 
| 1323 |                    Const("==>", _) $ _ $ B2) = strip_assums2 (B1,B2)
 | |
| 0 | 1324 |   | strip_assums2 (Const("all",_)$Abs(a,T,t1),
 | 
| 250 | 1325 |                    Const("all",_)$Abs(_,_,t2)) =
 | 
| 0 | 1326 | let val (B1,B2) = strip_assums2 (t1,t2) | 
| 1327 | in (Abs(a,T,B1), Abs(a,T,B2)) end | |
| 1328 | | strip_assums2 BB = BB; | |
| 1329 | ||
| 1330 | ||
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1331 | (*Faster normalization: skip assumptions that were lifted over*) | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1332 | fun norm_term_skip env 0 t = Envir.norm_term env t | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1333 |   | norm_term_skip env n (Const("all",_)$Abs(a,T,t)) =
 | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1334 |         let val Envir.Envir{iTs, ...} = env
 | 
| 1238 | 1335 | val T' = typ_subst_TVars iTs T | 
| 1336 | (*Must instantiate types of parameters because they are flattened; | |
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1337 | this could be a NEW parameter*) | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1338 | in all T' $ Abs(a, T', norm_term_skip env n t) end | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1339 |   | norm_term_skip env n (Const("==>", _) $ A $ B) =
 | 
| 1238 | 1340 | implies $ A $ norm_term_skip env (n-1) B | 
| 721 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1341 | | norm_term_skip env n t = error"norm_term_skip: too few assumptions??"; | 
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1342 | |
| 
479832ff2d29
Pure/thm/norm_term_skip: new, for skipping normalization of the context
 lcp parents: 
678diff
changeset | 1343 | |
| 0 | 1344 | (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) | 
| 250 | 1345 | Unifies B with Bi, replacing subgoal i (1 <= i <= n) | 
| 0 | 1346 | If match then forbid instantiations in proof state | 
| 1347 | If lifted then shorten the dpair using strip_assums2. | |
| 1348 | If eres_flg then simultaneously proves A1 by assumption. | |
| 250 | 1349 | nsubgoal is the number of new subgoals (written m above). | 
| 0 | 1350 | Curried so that resolution calls dest_state only once. | 
| 1351 | *) | |
| 4270 | 1352 | local exception COMPOSE | 
| 0 | 1353 | in | 
| 250 | 1354 | fun bicompose_aux match (state, (stpairs, Bs, Bi, C), lifted) | 
| 0 | 1355 | (eres_flg, orule, nsubgoal) = | 
| 1529 | 1356 |  let val Thm{der=sder, maxidx=smax, shyps=sshyps, hyps=shyps, ...} = state
 | 
| 1357 |      and Thm{der=rder, maxidx=rmax, shyps=rshyps, hyps=rhyps, 
 | |
| 2386 | 1358 | prop=rprop,...} = orule | 
| 1529 | 1359 | (*How many hyps to skip over during normalization*) | 
| 1238 | 1360 | and nlift = Logic.count_prems(strip_all_body Bi, | 
| 1361 | if eres_flg then ~1 else 0) | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1362 | val sign_ref = merge_thm_sgs(state,orule); | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1363 | val sign = Sign.deref sign_ref; | 
| 0 | 1364 | (** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **) | 
| 250 | 1365 |      fun addth As ((env as Envir.Envir {maxidx, ...}, tpairs), thq) =
 | 
| 1366 | let val normt = Envir.norm_term env; | |
| 1367 | (*perform minimal copying here by examining env*) | |
| 1368 | val normp = | |
| 1369 | if Envir.is_empty env then (tpairs, Bs @ As, C) | |
| 1370 | else | |
| 1371 | let val ntps = map (pairself normt) tpairs | |
| 2147 | 1372 | in if Envir.above (smax, env) then | 
| 1238 | 1373 | (*no assignments in state; normalize the rule only*) | 
| 1374 | if lifted | |
| 1375 | then (ntps, Bs @ map (norm_term_skip env nlift) As, C) | |
| 1376 | else (ntps, Bs @ map normt As, C) | |
| 1529 | 1377 | else if match then raise COMPOSE | 
| 250 | 1378 | else (*normalize the new rule fully*) | 
| 1379 | (ntps, map normt (Bs @ As), normt C) | |
| 1380 | end | |
| 1258 | 1381 | val th = (*tuned fix_shyps*) | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1382 |              Thm{sign_ref = sign_ref,
 | 
| 2386 | 1383 | der = infer_derivs (Bicompose(match, eres_flg, | 
| 1384 | 1 + length Bs, nsubgoal, env), | |
| 1385 | [rder,sder]), | |
| 1386 | maxidx = maxidx, | |
| 1387 | shyps = add_env_sorts (env, union_sort(rshyps,sshyps)), | |
| 1388 | hyps = union_term(rhyps,shyps), | |
| 1389 | prop = Logic.rule_of normp} | |
| 4270 | 1390 | in Seq.cons(th, thq) end handle COMPOSE => thq | 
| 0 | 1391 | val (rtpairs,rhorn) = Logic.strip_flexpairs(rprop); | 
| 1392 | val (rAs,B) = Logic.strip_prems(nsubgoal, [], rhorn) | |
| 1393 |        handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]);
 | |
| 1394 | (*Modify assumptions, deleting n-th if n>0 for e-resolution*) | |
| 1395 | fun newAs(As0, n, dpairs, tpairs) = | |
| 1396 | let val As1 = if !Logic.auto_rename orelse not lifted then As0 | |
| 250 | 1397 | else map (rename_bvars(dpairs,tpairs,B)) As0 | 
| 0 | 1398 | in (map (Logic.flatten_params n) As1) | 
| 250 | 1399 | handle TERM _ => | 
| 1400 |           raise THM("bicompose: 1st premise", 0, [orule])
 | |
| 0 | 1401 | end; | 
| 2147 | 1402 | val env = Envir.empty(Int.max(rmax,smax)); | 
| 0 | 1403 | val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); | 
| 1404 | val dpairs = BBi :: (rtpairs@stpairs); | |
| 1405 | (*elim-resolution: try each assumption in turn. Initially n=1*) | |
| 4270 | 1406 | fun tryasms (_, _, []) = Seq.empty | 
| 0 | 1407 | | tryasms (As, n, (t,u)::apairs) = | 
| 4270 | 1408 | (case Seq.pull(Unify.unifiers(sign, env, (t,u)::dpairs)) of | 
| 250 | 1409 | None => tryasms (As, n+1, apairs) | 
| 1410 | | cell as Some((_,tpairs),_) => | |
| 4270 | 1411 | Seq.it_right (addth (newAs(As, n, [BBi,(u,t)], tpairs))) | 
| 1412 | (Seq.make (fn()=> cell), | |
| 1413 | Seq.make (fn()=> Seq.pull (tryasms (As, n+1, apairs))))); | |
| 0 | 1414 |      fun eres [] = raise THM("bicompose: no premises", 0, [orule,state])
 | 
| 1415 | | eres (A1::As) = tryasms (As, 1, Logic.assum_pairs A1); | |
| 1416 | (*ordinary resolution*) | |
| 4270 | 1417 | fun res(None) = Seq.empty | 
| 250 | 1418 | | res(cell as Some((_,tpairs),_)) = | 
| 4270 | 1419 | Seq.it_right (addth(newAs(rev rAs, 0, [BBi], tpairs))) | 
| 1420 | (Seq.make (fn()=> cell), Seq.empty) | |
| 0 | 1421 | in if eres_flg then eres(rev rAs) | 
| 4270 | 1422 | else res(Seq.pull(Unify.unifiers(sign, env, dpairs))) | 
| 0 | 1423 | end; | 
| 1424 | end; (*open Sequence*) | |
| 1425 | ||
| 1426 | ||
| 1427 | fun bicompose match arg i state = | |
| 1428 | bicompose_aux match (state, dest_state(state,i), false) arg; | |
| 1429 | ||
| 1430 | (*Quick test whether rule is resolvable with the subgoal with hyps Hs | |
| 1431 | and conclusion B. If eres_flg then checks 1st premise of rule also*) | |
| 1432 | fun could_bires (Hs, B, eres_flg, rule) = | |
| 1433 | let fun could_reshyp (A1::_) = exists (apl(A1,could_unify)) Hs | |
| 250 | 1434 | | could_reshyp [] = false; (*no premise -- illegal*) | 
| 1435 | in could_unify(concl_of rule, B) andalso | |
| 1436 | (not eres_flg orelse could_reshyp (prems_of rule)) | |
| 0 | 1437 | end; | 
| 1438 | ||
| 1439 | (*Bi-resolution of a state with a list of (flag,rule) pairs. | |
| 1440 | Puts the rule above: rule/state. Renames vars in the rules. *) | |
| 250 | 1441 | fun biresolution match brules i state = | 
| 0 | 1442 | let val lift = lift_rule(state, i); | 
| 250 | 1443 | val (stpairs, Bs, Bi, C) = dest_state(state,i) | 
| 1444 | val B = Logic.strip_assums_concl Bi; | |
| 1445 | val Hs = Logic.strip_assums_hyp Bi; | |
| 1446 | val comp = bicompose_aux match (state, (stpairs, Bs, Bi, C), true); | |
| 4270 | 1447 | fun res [] = Seq.empty | 
| 250 | 1448 | | res ((eres_flg, rule)::brules) = | 
| 1449 | if could_bires (Hs, B, eres_flg, rule) | |
| 4270 | 1450 | then Seq.make (*delay processing remainder till needed*) | 
| 250 | 1451 | (fn()=> Some(comp (eres_flg, lift rule, nprems_of rule), | 
| 1452 | res brules)) | |
| 1453 | else res brules | |
| 4270 | 1454 | in Seq.flat (res brules) end; | 
| 0 | 1455 | |
| 1456 | ||
| 1457 | ||
| 2509 | 1458 | (*** Meta Simplification ***) | 
| 0 | 1459 | |
| 2509 | 1460 | (** diagnostics **) | 
| 0 | 1461 | |
| 1462 | exception SIMPLIFIER of string * thm; | |
| 1463 | ||
| 4045 | 1464 | fun prnt warn a = if warn then warning a else writeln a; | 
| 1465 | ||
| 1466 | fun prtm warn a sign t = | |
| 1467 | (prnt warn a; prnt warn (Sign.string_of_term sign t)); | |
| 1580 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 berghofe parents: 
1576diff
changeset | 1468 | |
| 4679 | 1469 | fun prthm warn a (thm as Thm{sign_ref, prop, ...}) =
 | 
| 1470 | (prtm warn a (Sign.deref sign_ref) prop); | |
| 1471 | ||
| 209 | 1472 | val trace_simp = ref false; | 
| 1473 | ||
| 4045 | 1474 | fun trace warn a = if !trace_simp then prnt warn a else (); | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1475 | |
| 4045 | 1476 | fun trace_term warn a sign t = | 
| 1477 | if !trace_simp then prtm warn a sign t else (); | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1478 | |
| 4045 | 1479 | fun trace_thm warn a (thm as Thm{sign_ref, prop, ...}) =
 | 
| 1480 | (trace_term warn a (Sign.deref sign_ref) prop); | |
| 209 | 1481 | |
| 1482 | ||
| 1580 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 berghofe parents: 
1576diff
changeset | 1483 | |
| 2509 | 1484 | (** meta simp sets **) | 
| 1485 | ||
| 1486 | (* basic components *) | |
| 1580 
e3fd931e6095
Added some functions which allow redirection of Isabelle's output
 berghofe parents: 
1576diff
changeset | 1487 | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1488 | type rrule = {thm: thm, lhs: term, elhs: term, fo: bool, perm: bool};
 | 
| 2509 | 1489 | type cong = {thm: thm, lhs: term};
 | 
| 3577 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 1490 | type simproc = | 
| 
9715b6e3ec5f
added prems argument to simplification procedures;
 wenzelm parents: 
3565diff
changeset | 1491 |  {name: string, proc: Sign.sg -> thm list -> term -> thm option, lhs: cterm, id: stamp};
 | 
| 288 
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
 nipkow parents: 
274diff
changeset | 1492 | |
| 3550 | 1493 | fun eq_rrule ({thm = Thm {prop = p1, ...}, ...}: rrule,
 | 
| 2509 | 1494 |   {thm = Thm {prop = p2, ...}, ...}: rrule) = p1 aconv p2;
 | 
| 1495 | ||
| 3550 | 1496 | fun eq_cong ({thm = Thm {prop = p1, ...}, ...}: cong,
 | 
| 1497 |   {thm = Thm {prop = p2, ...}, ...}: cong) = p1 aconv p2;
 | |
| 1498 | ||
| 1499 | fun eq_prem (Thm {prop = p1, ...}, Thm {prop = p2, ...}) = p1 aconv p2;
 | |
| 1500 | ||
| 1501 | fun eq_simproc ({id = s1, ...}:simproc, {id = s2, ...}:simproc) = (s1 = s2);
 | |
| 1502 | ||
| 1503 | fun mk_simproc (name, proc, lhs, id) = | |
| 1504 |   {name = name, proc = proc, lhs = lhs, id = id};
 | |
| 2509 | 1505 | |
| 1506 | ||
| 1507 | (* datatype mss *) | |
| 288 
b00ce6a1fe27
Implemented "ordered rewriting": rules which merely permute variables, such
 nipkow parents: 
274diff
changeset | 1508 | |
| 2509 | 1509 | (* | 
| 1510 | A "mss" contains data needed during conversion: | |
| 1511 | rules: discrimination net of rewrite rules; | |
| 1512 | congs: association list of congruence rules; | |
| 1513 | procs: discrimination net of simplification procedures | |
| 1514 | (functions that prove rewrite rules on the fly); | |
| 1515 | bounds: names of bound variables already used | |
| 1516 | (for generating new names when rewriting under lambda abstractions); | |
| 1517 | prems: current premises; | |
| 4679 | 1518 | mk_rews: mk: turns simplification thms into rewrite rules; | 
| 1519 | mk_sym: turns == around; (needs Drule!) | |
| 1520 | mk_eq_True: turns P into P == True - logic specific; | |
| 2509 | 1521 | termless: relation for ordered rewriting; | 
| 1028 | 1522 | *) | 
| 0 | 1523 | |
| 2509 | 1524 | datatype meta_simpset = | 
| 1525 |   Mss of {
 | |
| 1526 | rules: rrule Net.net, | |
| 1527 | congs: (string * cong) list, | |
| 1528 | procs: simproc Net.net, | |
| 1529 | bounds: string list, | |
| 1530 | prems: thm list, | |
| 4679 | 1531 |     mk_rews: {mk: thm -> thm list,
 | 
| 1532 | mk_sym: thm -> thm option, | |
| 1533 | mk_eq_True: thm -> thm option}, | |
| 2509 | 1534 | termless: term * term -> bool}; | 
| 1535 | ||
| 1536 | fun mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless) = | |
| 1537 |   Mss {rules = rules, congs = congs, procs = procs, bounds = bounds,
 | |
| 4679 | 1538 | prems=prems, mk_rews=mk_rews, termless=termless}; | 
| 1539 | ||
| 1540 | fun upd_rules(Mss{rules,congs,procs,bounds,prems,mk_rews,termless}, rules') =
 | |
| 1541 | mk_mss(rules',congs,procs,bounds,prems,mk_rews,termless); | |
| 2509 | 1542 | |
| 1543 | val empty_mss = | |
| 4679 | 1544 |   let val mk_rews = {mk = K [], mk_sym = K None, mk_eq_True = K None}
 | 
| 1545 | in mk_mss (Net.empty, [], Net.empty, [], [], mk_rews, Term.termless) end; | |
| 2509 | 1546 | |
| 1547 | ||
| 1548 | ||
| 1549 | (** simpset operations **) | |
| 1550 | ||
| 3550 | 1551 | (* dest_mss *) | 
| 1552 | ||
| 1553 | fun dest_mss (Mss {rules, congs, procs, ...}) =
 | |
| 1554 |   {simps = map (fn (_, {thm, ...}) => thm) (Net.dest rules),
 | |
| 1555 |    congs = map (fn (_, {thm, ...}) => thm) congs,
 | |
| 1556 | procs = | |
| 1557 |      map (fn (_, {name, lhs, id, ...}) => ((name, lhs), id)) (Net.dest procs)
 | |
| 1558 | |> partition_eq eq_snd | |
| 1559 | |> map (fn ps => (#1 (#1 (hd ps)), map (#2 o #1) ps))}; | |
| 1560 | ||
| 1561 | ||
| 1562 | (* merge_mss *) (*NOTE: ignores mk_rews and termless of 2nd mss*) | |
| 1563 | ||
| 1564 | fun merge_mss | |
| 1565 |  (Mss {rules = rules1, congs = congs1, procs = procs1, bounds = bounds1,
 | |
| 1566 | prems = prems1, mk_rews, termless}, | |
| 1567 |   Mss {rules = rules2, congs = congs2, procs = procs2, bounds = bounds2,
 | |
| 1568 | prems = prems2, ...}) = | |
| 1569 | mk_mss | |
| 1570 | (Net.merge (rules1, rules2, eq_rrule), | |
| 1571 | generic_merge (eq_cong o pairself snd) I I congs1 congs2, | |
| 1572 | Net.merge (procs1, procs2, eq_simproc), | |
| 1573 | merge_lists bounds1 bounds2, | |
| 1574 | generic_merge eq_prem I I prems1 prems2, | |
| 1575 | mk_rews, termless); | |
| 1576 | ||
| 4679 | 1577 | (* add_simps *) | 
| 3550 | 1578 | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1579 | fun mk_rrule2{thm,lhs,perm} =
 | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1580 | let val elhs = Pattern.eta_contract lhs | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1581 | val fo = Pattern.first_order elhs orelse not(Pattern.pattern elhs) | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1582 |   in {thm=thm,lhs=lhs,elhs=elhs,fo=fo,perm=perm} end
 | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1583 | |
| 4679 | 1584 | fun insert_rrule(mss as Mss {rules,...},
 | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1585 |                  rrule as {thm,lhs,perm}) =
 | 
| 4679 | 1586 | (trace_thm false "Adding rewrite rule:" thm; | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1587 |    let val rrule2 as {elhs,...} = mk_rrule2 rrule
 | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1588 | val rules' = Net.insert_term ((elhs, rrule2), rules, eq_rrule) | 
| 4679 | 1589 | in upd_rules(mss,rules') end | 
| 1590 | handle Net.INSERT => | |
| 4785 | 1591 | (prthm true "Ignoring duplicate rewrite rule:" thm; mss)); | 
| 4679 | 1592 | |
| 1593 | fun vperm (Var _, Var _) = true | |
| 1594 | | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) | |
| 1595 | | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) | |
| 1596 | | vperm (t, u) = (t = u); | |
| 1597 | ||
| 1598 | fun var_perm (t, u) = | |
| 1599 | vperm (t, u) andalso eq_set_term (term_vars t, term_vars u); | |
| 1600 | ||
| 1601 | (* FIXME: it seems that the conditions on extra variables are too liberal if | |
| 1602 | prems are nonempty: does solving the prems really guarantee instantiation of | |
| 1603 | all its Vars? Better: a dynamic check each time a rule is applied. | |
| 1604 | *) | |
| 1605 | fun rewrite_rule_extra_vars prems elhs erhs = | |
| 1606 | not ((term_vars erhs) subset | |
| 1607 | (union_term (term_vars elhs, List.concat(map term_vars prems)))) | |
| 1608 | orelse | |
| 1609 | not ((term_tvars erhs) subset | |
| 1610 | (term_tvars elhs union List.concat(map term_tvars prems))); | |
| 2509 | 1611 | |
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1612 | (*Simple test for looping rewrite rules and stupid orientations*) | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1613 | fun reorient sign prems lhs rhs = | 
| 4679 | 1614 | rewrite_rule_extra_vars prems lhs rhs | 
| 1615 | orelse | |
| 1616 | is_Var (head_of lhs) | |
| 1617 | orelse | |
| 4684 | 1618 | (exists (apl (lhs, Logic.occs)) (rhs :: prems)) | 
| 4679 | 1619 | orelse | 
| 1620 | (null prems andalso | |
| 1621 | Pattern.matches (#tsig (Sign.rep_sg sign)) (lhs, rhs)) | |
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1622 | (*the condition "null prems" is necessary because conditional rewrites | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1623 | with extra variables in the conditions may terminate although | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1624 | the rhs is an instance of the lhs. Example: ?m < ?n ==> f(?n) == f(?m)*) | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1625 | orelse | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1626 | (is_Const lhs andalso not(is_Const rhs)) | 
| 4679 | 1627 | |
| 1628 | fun decomp_simp(thm as Thm {sign_ref, prop, ...}) =
 | |
| 1629 | let val sign = Sign.deref sign_ref; | |
| 1630 | val prems = Logic.strip_imp_prems prop; | |
| 1631 | val concl = Logic.strip_imp_concl prop; | |
| 1632 | val (lhs, rhs) = Logic.dest_equals concl handle TERM _ => | |
| 1633 |         raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm)
 | |
| 1634 | val elhs = Pattern.eta_contract lhs; | |
| 1635 | val erhs = Pattern.eta_contract rhs; | |
| 1636 | val perm = var_perm (elhs, erhs) andalso not (elhs aconv erhs) | |
| 1637 | andalso not (is_Var elhs) | |
| 1638 | in (sign,prems,lhs,rhs,perm) end; | |
| 1639 | ||
| 1640 | fun mk_eq_True (Mss{mk_rews={mk_eq_True,...},...}) thm =
 | |
| 4713 | 1641 | case mk_eq_True thm of | 
| 1642 | None => [] | |
| 1643 | | Some eq_True => let val (_,_,lhs,_,_) = decomp_simp eq_True | |
| 1644 |                     in [{thm=eq_True, lhs=lhs, perm=false}] end;
 | |
| 1645 | ||
| 1646 | (* create the rewrite rule and possibly also the ==True variant, | |
| 1647 | in case there are extra vars on the rhs *) | |
| 1648 | fun rrule_eq_True(thm,lhs,rhs,mss,thm2) = | |
| 1649 |   let val rrule = {thm=thm, lhs=lhs, perm=false}
 | |
| 1650 | in if (term_vars rhs) subset (term_vars lhs) andalso | |
| 1651 | (term_tvars rhs) subset (term_tvars lhs) | |
| 1652 | then [rrule] | |
| 1653 | else mk_eq_True mss thm2 @ [rrule] | |
| 1654 | end; | |
| 4679 | 1655 | |
| 1656 | fun mk_rrule mss thm = | |
| 1657 | let val (_,prems,lhs,rhs,perm) = decomp_simp thm | |
| 4713 | 1658 |   in if perm then [{thm=thm, lhs=lhs, perm=true}] else
 | 
| 4679 | 1659 | (* weak test for loops: *) | 
| 1660 | if rewrite_rule_extra_vars prems lhs rhs orelse | |
| 1661 | is_Var (head_of lhs) (* mk_cases may do this! *) | |
| 1662 | then mk_eq_True mss thm | |
| 4713 | 1663 | else rrule_eq_True(thm,lhs,rhs,mss,thm) | 
| 0 | 1664 | end; | 
| 1665 | ||
| 4679 | 1666 | fun orient_rrule mss thm = | 
| 1667 | let val (sign,prems,lhs,rhs,perm) = decomp_simp thm | |
| 4713 | 1668 |   in if perm then [{thm=thm,lhs=lhs,perm=true}]
 | 
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1669 | else if reorient sign prems lhs rhs | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 1670 | then if reorient sign prems rhs lhs | 
| 4679 | 1671 | then mk_eq_True mss thm | 
| 1672 |                else let val Mss{mk_rews={mk_sym,...},...} = mss
 | |
| 4713 | 1673 | in case mk_sym thm of | 
| 1674 | None => [] | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1675 | | Some thm' => | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1676 | let val (_,_,lhs',rhs',_) = decomp_simp thm' | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1677 | in rrule_eq_True(thm',lhs',rhs',mss,thm) end | 
| 4679 | 1678 | end | 
| 4713 | 1679 | else rrule_eq_True(thm,lhs,rhs,mss,thm) | 
| 4679 | 1680 | end; | 
| 2509 | 1681 | |
| 4679 | 1682 | fun extract_rews(Mss{mk_rews = {mk,...},...},thms) = flat(map mk thms);
 | 
| 87 | 1683 | |
| 4679 | 1684 | fun orient_comb_simps comb mk_rrule (mss,thms) = | 
| 1685 | let val rews = extract_rews(mss,thms) | |
| 4713 | 1686 | val rrules = flat (map mk_rrule rews) | 
| 4679 | 1687 | in foldl comb (mss,rrules) end | 
| 4667 
6328d427a339
Tried to reorganize rewriter a little. More to be done.
 nipkow parents: 
4589diff
changeset | 1688 | |
| 4679 | 1689 | (* Add rewrite rules explicitly; do not reorient! *) | 
| 1690 | fun add_simps(mss,thms) = | |
| 1691 | orient_comb_simps insert_rrule (mk_rrule mss) (mss,thms); | |
| 0 | 1692 | |
| 4679 | 1693 | fun mss_of thms = | 
| 4713 | 1694 | foldl insert_rrule (empty_mss, flat(map (mk_rrule empty_mss) thms)); | 
| 2509 | 1695 | |
| 4713 | 1696 | fun extract_safe_rrules(mss,thm) = | 
| 1697 | flat (map (orient_rrule mss) (extract_rews(mss,[thm]))); | |
| 2509 | 1698 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 1699 | fun add_safe_simp(mss,thm) = | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 1700 | foldl insert_rrule (mss, extract_safe_rrules(mss,thm)) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 1701 | |
| 2509 | 1702 | (* del_simps *) | 
| 1703 | ||
| 4679 | 1704 | fun del_rrule(mss as Mss {rules,...},
 | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1705 |               rrule as {thm, elhs, ...}) =
 | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1706 | (upd_rules(mss, Net.delete_term ((elhs, rrule), rules, eq_rrule)) | 
| 4679 | 1707 | handle Net.DELETE => | 
| 4785 | 1708 | (prthm true "Rewrite rule not in simpset:" thm; mss)); | 
| 4667 
6328d427a339
Tried to reorganize rewriter a little. More to be done.
 nipkow parents: 
4589diff
changeset | 1709 | |
| 4679 | 1710 | fun del_simps(mss,thms) = | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1711 | orient_comb_simps del_rrule (map mk_rrule2 o mk_rrule mss) (mss,thms); | 
| 0 | 1712 | |
| 2509 | 1713 | |
| 2626 | 1714 | (* add_congs *) | 
| 0 | 1715 | |
| 4679 | 1716 | fun add_cong (Mss {rules,congs,procs,bounds,prems,mk_rews,termless}, thm) =
 | 
| 2509 | 1717 | let | 
| 1718 | val (lhs, _) = Logic.dest_equals (concl_of thm) handle TERM _ => | |
| 1719 |       raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | |
| 1720 | (* val lhs = Pattern.eta_contract lhs; *) | |
| 1721 | val (a, _) = dest_Const (head_of lhs) handle TERM _ => | |
| 1722 |       raise SIMPLIFIER ("Congruence must start with a constant", thm);
 | |
| 1723 | in | |
| 1724 |     mk_mss (rules, (a, {lhs = lhs, thm = thm}) :: congs, procs, bounds,
 | |
| 1725 | prems, mk_rews, termless) | |
| 0 | 1726 | end; | 
| 1727 | ||
| 1728 | val (op add_congs) = foldl add_cong; | |
| 1729 | ||
| 2509 | 1730 | |
| 2626 | 1731 | (* del_congs *) | 
| 1732 | ||
| 4679 | 1733 | fun del_cong (Mss {rules,congs,procs,bounds,prems,mk_rews,termless}, thm) =
 | 
| 2626 | 1734 | let | 
| 1735 | val (lhs, _) = Logic.dest_equals (concl_of thm) handle TERM _ => | |
| 1736 |       raise SIMPLIFIER ("Congruence not a meta-equality", thm);
 | |
| 1737 | (* val lhs = Pattern.eta_contract lhs; *) | |
| 1738 | val (a, _) = dest_Const (head_of lhs) handle TERM _ => | |
| 1739 |       raise SIMPLIFIER ("Congruence must start with a constant", thm);
 | |
| 1740 | in | |
| 1741 | mk_mss (rules, filter (fn (x,_)=> x<>a) congs, procs, bounds, | |
| 1742 | prems, mk_rews, termless) | |
| 1743 | end; | |
| 1744 | ||
| 1745 | val (op del_congs) = foldl del_cong; | |
| 1746 | ||
| 1747 | ||
| 2509 | 1748 | (* add_simprocs *) | 
| 1749 | ||
| 4679 | 1750 | fun add_proc (mss as Mss {rules,congs,procs,bounds,prems,mk_rews,termless},
 | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1751 |     (name, lhs as Cterm {sign_ref, t, ...}, proc, id)) =
 | 
| 4045 | 1752 |   (trace_term false ("Adding simplification procedure " ^ quote name ^ " for:")
 | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1753 | (Sign.deref sign_ref) t; | 
| 2509 | 1754 | mk_mss (rules, congs, | 
| 3550 | 1755 | Net.insert_term ((t, mk_simproc (name, proc, lhs, id)), procs, eq_simproc) | 
| 4785 | 1756 | handle Net.INSERT => (trace true "Ignored duplicate"; procs), | 
| 2509 | 1757 | bounds, prems, mk_rews, termless)); | 
| 0 | 1758 | |
| 3550 | 1759 | fun add_simproc (mss, (name, lhss, proc, id)) = | 
| 1760 | foldl add_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss); | |
| 1761 | ||
| 2509 | 1762 | val add_simprocs = foldl add_simproc; | 
| 1763 | ||
| 1764 | ||
| 1765 | (* del_simprocs *) | |
| 0 | 1766 | |
| 4679 | 1767 | fun del_proc (mss as Mss {rules,congs,procs,bounds,prems,mk_rews,termless},
 | 
| 3550 | 1768 |     (name, lhs as Cterm {t, ...}, proc, id)) =
 | 
| 2509 | 1769 | mk_mss (rules, congs, | 
| 3550 | 1770 | Net.delete_term ((t, mk_simproc (name, proc, lhs, id)), procs, eq_simproc) | 
| 4785 | 1771 | handle Net.DELETE => (trace true "Simplification procedure not in simpset"; procs), | 
| 3550 | 1772 | bounds, prems, mk_rews, termless); | 
| 1773 | ||
| 1774 | fun del_simproc (mss, (name, lhss, proc, id)) = | |
| 1775 | foldl del_proc (mss, map (fn lhs => (name, lhs, proc, id)) lhss); | |
| 2509 | 1776 | |
| 1777 | val del_simprocs = foldl del_simproc; | |
| 0 | 1778 | |
| 1779 | ||
| 2509 | 1780 | (* prems *) | 
| 1781 | ||
| 4679 | 1782 | fun add_prems (Mss {rules,congs,procs,bounds,prems,mk_rews,termless}, thms) =
 | 
| 2509 | 1783 | mk_mss (rules, congs, procs, bounds, thms @ prems, mk_rews, termless); | 
| 1784 | ||
| 1785 | fun prems_of_mss (Mss {prems, ...}) = prems;
 | |
| 1786 | ||
| 1787 | ||
| 1788 | (* mk_rews *) | |
| 1789 | ||
| 1790 | fun set_mk_rews | |
| 4679 | 1791 |   (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, mk) =
 | 
| 1792 | mk_mss (rules, congs, procs, bounds, prems, | |
| 1793 |             {mk=mk, mk_sym= #mk_sym mk_rews, mk_eq_True= #mk_eq_True mk_rews},
 | |
| 1794 | termless); | |
| 2509 | 1795 | |
| 4679 | 1796 | fun set_mk_sym | 
| 1797 |   (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, mk_sym) =
 | |
| 1798 | mk_mss (rules, congs, procs, bounds, prems, | |
| 1799 |             {mk= #mk mk_rews, mk_sym= mk_sym, mk_eq_True= #mk_eq_True mk_rews},
 | |
| 1800 | termless); | |
| 2509 | 1801 | |
| 4679 | 1802 | fun set_mk_eq_True | 
| 1803 |   (Mss {rules, congs, procs, bounds, prems, mk_rews, termless}, mk_eq_True) =
 | |
| 1804 | mk_mss (rules, congs, procs, bounds, prems, | |
| 1805 |             {mk= #mk mk_rews, mk_sym= #mk_sym mk_rews, mk_eq_True= mk_eq_True},
 | |
| 1806 | termless); | |
| 2509 | 1807 | |
| 1808 | (* termless *) | |
| 1809 | ||
| 1810 | fun set_termless | |
| 1811 |   (Mss {rules, congs, procs, bounds, prems, mk_rews, termless = _}, termless) =
 | |
| 1812 | mk_mss (rules, congs, procs, bounds, prems, mk_rews, termless); | |
| 1813 | ||
| 1814 | ||
| 1815 | ||
| 1816 | (** rewriting **) | |
| 1817 | ||
| 1818 | (* | |
| 1819 | Uses conversions, omitting proofs for efficiency. See: | |
| 1820 | L C Paulson, A higher-order implementation of rewriting, | |
| 1821 | Science of Computer Programming 3 (1983), pages 119-149. | |
| 1822 | *) | |
| 0 | 1823 | |
| 1824 | type prover = meta_simpset -> thm -> thm option; | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1825 | type termrec = (Sign.sg_ref * term list) * term; | 
| 0 | 1826 | type conv = meta_simpset -> termrec -> termrec; | 
| 1827 | ||
| 4116 | 1828 | fun check_conv (thm as Thm{shyps,hyps,prop,sign_ref,der,...}, prop0, ders) =
 | 
| 4045 | 1829 | let fun err() = (trace_thm false "Proved wrong thm (Check subgoaler?)" thm; | 
| 4785 | 1830 | trace_term false "Should have proved:" (Sign.deref sign_ref) prop0; | 
| 432 | 1831 | None) | 
| 0 | 1832 | val (lhs0,_) = Logic.dest_equals(Logic.strip_imp_concl prop0) | 
| 1833 | in case prop of | |
| 1834 |        Const("==",_) $ lhs $ rhs =>
 | |
| 1835 | if (lhs = lhs0) orelse | |
| 427 
4ce2ce940faf
ordered rewriting applies to conditional rules as well now
 nipkow parents: 
421diff
changeset | 1836 | (lhs aconv Envir.norm_term (Envir.empty 0) lhs0) | 
| 4045 | 1837 | then (trace_thm false "SUCCEEDED" thm; | 
| 4713 | 1838 | Some(rhs, (shyps, hyps, der::ders))) | 
| 0 | 1839 | else err() | 
| 1840 | | _ => err() | |
| 1841 | end; | |
| 1842 | ||
| 659 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1843 | fun ren_inst(insts,prop,pat,obj) = | 
| 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1844 | let val ren = match_bvs(pat,obj,[]) | 
| 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1845 | fun renAbs(Abs(x,T,b)) = | 
| 1576 
af8f43f742a0
Added some optimized versions of functions dealing with sets
 berghofe parents: 
1569diff
changeset | 1846 | Abs(case assoc_string(ren,x) of None => x | Some(y) => y, T, renAbs(b)) | 
| 659 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1847 | | renAbs(f$t) = renAbs(f) $ renAbs(t) | 
| 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1848 | | renAbs(t) = t | 
| 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1849 | in subst_vars insts (if null(ren) then prop else renAbs(prop)) end; | 
| 678 
6151b7f3b606
Modified pattern.ML to perform proper matching of Higher-Order Patterns.
 nipkow parents: 
659diff
changeset | 1850 | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1851 | fun incr_insts i (in1:(indexname*typ)list,in2:(indexname*term)list) = | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1852 | let fun incr ((a,n),x) = ((a,n+i),x) | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1853 | in (map incr in1, map incr in2) end; | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1854 | |
| 1258 | 1855 | fun add_insts_sorts ((iTs, is), Ss) = | 
| 1856 | add_typs_sorts (map snd iTs, add_terms_sorts (map snd is, Ss)); | |
| 1857 | ||
| 659 
95ed2ccb4438
Prepared the code for preservation of bound var names during rewriting. Still
 nipkow parents: 
623diff
changeset | 1858 | |
| 2509 | 1859 | (* mk_procrule *) | 
| 1860 | ||
| 4679 | 1861 | fun mk_procrule thm = | 
| 1862 | let val (_,prems,lhs,rhs,_) = decomp_simp thm | |
| 1863 | in if rewrite_rule_extra_vars prems lhs rhs | |
| 4785 | 1864 | then (prthm true "Extra vars on rhs:" thm; []) | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1865 |      else [mk_rrule2{thm = thm, lhs = lhs, perm = false}]
 | 
| 2509 | 1866 | end; | 
| 1867 | ||
| 1868 | ||
| 1869 | (* conversion to apply the meta simpset to a term *) | |
| 1870 | ||
| 1871 | (* | |
| 1872 | we try in order: | |
| 1873 | (1) beta reduction | |
| 1874 | (2) unconditional rewrite rules | |
| 1875 | (3) conditional rewrite rules | |
| 3550 | 1876 | (4) simplification procedures | 
| 4116 | 1877 | |
| 1878 | IMPORTANT: rewrite rules must not introduce new Vars or TVars! | |
| 1879 | ||
| 2509 | 1880 | *) | 
| 1881 | ||
| 4116 | 1882 | fun rewritec (prover,sign_reft,maxt) | 
| 4679 | 1883 |              (mss as Mss{rules, procs, termless, prems, ...}) 
 | 
| 4713 | 1884 | (t:term,etc as (shypst,hypst,ders)) = | 
| 3550 | 1885 | let | 
| 4713 | 1886 | val signt = Sign.deref sign_reft; | 
| 1887 | val tsigt = Sign.tsig_of signt; | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1888 |     fun rew{thm as Thm{sign_ref,der,shyps,hyps,prop,maxidx,...},
 | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1889 | lhs, elhs, fo, perm} = | 
| 4713 | 1890 | let | 
| 1891 | val _ = if Sign.subsig (Sign.deref sign_ref, signt) then () | |
| 4785 | 1892 | else (trace_thm true "Rewrite rule from different theory:" thm; | 
| 4713 | 1893 | raise Pattern.MATCH); | 
| 1894 | val rprop = if maxt = ~1 then prop | |
| 1895 | else Logic.incr_indexes([],maxt+1) prop; | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1896 | val insts = if fo then Pattern.first_order_match tsigt (elhs,t) | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1897 | else Pattern.match tsigt (elhs,t); | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1898 | val insts = if maxt = ~1 then insts else incr_insts (maxt+1) insts | 
| 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1899 | val prop' = ren_inst(insts,rprop,lhs,t); | 
| 4713 | 1900 | val hyps' = union_term(hyps,hypst); | 
| 1901 | val shyps' = add_insts_sorts (insts, union_sort(shyps,shypst)); | |
| 1902 | val unconditional = (Logic.count_prems(prop',0) = 0); | |
| 1903 | val maxidx' = if unconditional then maxt else maxidx+maxt+1 | |
| 1904 |         val ct' = Cterm{sign_ref = sign_reft,       (*used for deriv only*)
 | |
| 1905 | t = prop', T = propT, maxidx = maxidx'} | |
| 1906 | val der' = infer_derivs (RewriteC ct', [der]); | |
| 1907 |         val thm' = Thm{sign_ref = sign_reft, der = der', shyps = shyps',
 | |
| 1908 | hyps = hyps', prop = prop', maxidx = maxidx'} | |
| 1909 | val (lhs',rhs') = Logic.dest_equals(Logic.strip_imp_concl prop') | |
| 1910 | in | |
| 1911 | if perm andalso not(termless(rhs',lhs')) then None | |
| 1912 | else (trace_thm false "Applying instance of rewrite rule:" thm; | |
| 1913 | if unconditional | |
| 1914 | then (trace_thm false "Rewriting:" thm'; | |
| 1915 | Some(rhs', (shyps', hyps', der'::ders))) | |
| 1916 | else (trace_thm false "Trying to rewrite:" thm'; | |
| 1917 | case prover mss thm' of | |
| 1918 | None => (trace_thm false "FAILED" thm'; None) | |
| 1919 | | Some(thm2) => check_conv(thm2,prop',ders))) | |
| 1659 | 1920 | end | 
| 2509 | 1921 | |
| 4713 | 1922 | fun rews [] = None | 
| 1923 | | rews (rrule :: rrules) = | |
| 1924 | let val opt = rew rrule handle Pattern.MATCH => None | |
| 1925 | in case opt of None => rews rrules | some => some end; | |
| 1926 | ||
| 1927 | fun sort_rrules rrs = let | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1928 |       fun is_simple({thm as Thm{prop,...}, ...}:rrule) = case prop of 
 | 
| 4713 | 1929 |                                       Const("==",_) $ _ $ _ => true
 | 
| 1930 | | _ => false | |
| 1931 | fun sort [] (re1,re2) = re1 @ re2 | |
| 1932 | | sort (rr::rrs) (re1,re2) = if is_simple rr | |
| 1933 | then sort rrs (rr::re1,re2) | |
| 1934 | else sort rrs (re1,rr::re2) | |
| 1935 | in sort rrs ([],[]) end | |
| 1936 | ||
| 1937 | fun proc_rews _ ([]:simproc list) = None | |
| 1938 |       | proc_rews eta_t ({name, proc, lhs = Cterm {t = plhs, ...}, ...} :: ps) =
 | |
| 1939 | if Pattern.matches tsigt (plhs, t) then | |
| 1940 |             (trace_term false ("Trying procedure " ^ quote name ^ " on:") signt eta_t;
 | |
| 1941 | case proc signt prems eta_t of | |
| 1942 | None => (trace false "FAILED"; proc_rews eta_t ps) | |
| 1943 | | Some raw_thm => | |
| 4397 | 1944 |                  (trace_thm false ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm;
 | 
| 4713 | 1945 | (case rews (mk_procrule raw_thm) of | 
| 1946 | None => (trace false "IGNORED"; proc_rews eta_t ps) | |
| 1947 | | some => some))) | |
| 1948 | else proc_rews eta_t ps; | |
| 1949 | in case t of | |
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 1950 | Abs (_, _, body) $ u => Some (subst_bound (u, body), etc) | 
| 4713 | 1951 | | _ => (case rews (sort_rrules (Net.match_term rules t)) of | 
| 1952 | None => proc_rews (Pattern.eta_contract t) | |
| 1953 | (Net.match_term procs t) | |
| 1954 | | some => some) | |
| 0 | 1955 | end; | 
| 1956 | ||
| 2509 | 1957 | |
| 1958 | (* conversion to apply a congruence rule to a term *) | |
| 1959 | ||
| 4713 | 1960 | fun congc (prover,sign_reft,maxt) {thm=cong,lhs=lhs} (t,(shypst,hypst,ders)) =
 | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1961 | let val signt = Sign.deref sign_reft; | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1962 | val tsig = Sign.tsig_of signt; | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1963 |       val Thm{sign_ref,der,shyps,hyps,maxidx,prop,...} = cong
 | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1964 | val _ = if Sign.subsig(Sign.deref sign_ref,signt) then () | 
| 208 
342f88d2e8ab
optimized simplifier - signature of rewritten term stays constant
 nipkow parents: 
200diff
changeset | 1965 |                  else error("Congruence rule from different theory")
 | 
| 2147 | 1966 | val rprop = if maxt = ~1 then prop | 
| 1967 | else Logic.incr_indexes([],maxt+1) prop; | |
| 1968 | val rlhs = if maxt = ~1 then lhs | |
| 1065 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 nipkow parents: 
1028diff
changeset | 1969 | else fst(Logic.dest_equals(Logic.strip_imp_concl rprop)) | 
| 1569 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 nipkow parents: 
1566diff
changeset | 1970 | val insts = Pattern.match tsig (rlhs,t) | 
| 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 nipkow parents: 
1566diff
changeset | 1971 | (* Pattern.match can raise Pattern.MATCH; | 
| 
a89f246dee0e
Non-matching congruence rule in rewriter is simply ignored.
 nipkow parents: 
1566diff
changeset | 1972 | is handled when congc is called *) | 
| 1065 
8425cb5acb77
Fixed old bug in the simplifier. Term to be simplified now carries around its
 nipkow parents: 
1028diff
changeset | 1973 | val prop' = ren_inst(insts,rprop,rlhs,t); | 
| 2177 
8b365a3a6ed1
Changed some mem, ins and union calls to be monomorphic
 paulson parents: 
2163diff
changeset | 1974 | val shyps' = add_insts_sorts (insts, union_sort(shyps,shypst)) | 
| 1529 | 1975 | val maxidx' = maxidx_of_term prop' | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1976 |       val ct' = Cterm{sign_ref = sign_reft,     (*used for deriv only*)
 | 
| 2386 | 1977 | t = prop', | 
| 1978 | T = propT, | |
| 1979 | maxidx = maxidx'} | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 1980 |       val thm' = Thm{sign_ref = sign_reft, 
 | 
| 3550 | 1981 | der = infer_derivs (CongC ct', [der]), | 
| 2386 | 1982 | shyps = shyps', | 
| 1983 | hyps = union_term(hyps,hypst), | |
| 1529 | 1984 | prop = prop', | 
| 2386 | 1985 | maxidx = maxidx'}; | 
| 4785 | 1986 | val unit = trace_thm false "Applying congruence rule:" thm'; | 
| 112 
009ae5c85ae9
Changed the simplifier: if the subgoaler proves an unexpected thm, chances
 nipkow parents: 
87diff
changeset | 1987 |       fun err() = error("Failed congruence proof!")
 | 
| 0 | 1988 | |
| 1989 | in case prover thm' of | |
| 112 
009ae5c85ae9
Changed the simplifier: if the subgoaler proves an unexpected thm, chances
 nipkow parents: 
87diff
changeset | 1990 | None => err() | 
| 1529 | 1991 | | Some(thm2) => (case check_conv(thm2,prop',ders) of | 
| 405 
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
 nipkow parents: 
399diff
changeset | 1992 | None => err() | some => some) | 
| 0 | 1993 | end; | 
| 1994 | ||
| 4713 | 1995 | fun bottomc ((simprem,useprem,mutsimp),prover,sign_ref,maxidx) = | 
| 1996 | let | |
| 1997 | fun botc fail mss trec = | |
| 2386 | 1998 | (case subc mss trec of | 
| 1999 | some as Some(trec1) => | |
| 4116 | 2000 | (case rewritec (prover,sign_ref,maxidx) mss trec1 of | 
| 2386 | 2001 | Some(trec2) => botc false mss trec2 | 
| 2002 | | None => some) | |
| 2003 | | None => | |
| 4116 | 2004 | (case rewritec (prover,sign_ref,maxidx) mss trec of | 
| 2386 | 2005 | Some(trec2) => botc false mss trec2 | 
| 2006 | | None => if fail then None else Some(trec))) | |
| 0 | 2007 | |
| 4713 | 2008 | and try_botc mss trec = (case botc true mss trec of | 
| 2386 | 2009 | Some(trec1) => trec1 | 
| 2010 | | None => trec) | |
| 405 
c97514f63633
Internale optimization of the simplifier: in case a subterm stays unchanged,
 nipkow parents: 
399diff
changeset | 2011 | |
| 4713 | 2012 |     and subc (mss as Mss{rules,congs,procs,bounds,prems,mk_rews,termless})
 | 
| 2013 | (trec as (t0:term,etc:sort list*term list * rule mtree list)) = | |
| 1529 | 2014 | (case t0 of | 
| 2386 | 2015 | Abs(a,T,t) => | 
| 2016 | let val b = variant bounds a | |
| 2017 |                  val v = Free("." ^ b,T)
 | |
| 2509 | 2018 | val mss' = mk_mss (rules, congs, procs, b :: bounds, prems, mk_rews, termless) | 
| 4713 | 2019 | in case botc true mss' (subst_bound(v,t),etc) of | 
| 2020 | Some(t',etc') => Some(Abs(a, T, abstract_over(v,t')), etc') | |
| 2386 | 2021 | | None => None | 
| 2022 | end | |
| 2023 | | t$u => (case t of | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2024 |              Const("==>",_)$s  => Some(impc(s,u,mss,etc))
 | 
| 2386 | 2025 | | Abs(_,_,body) => | 
| 4713 | 2026 | let val trec = (subst_bound(u,body), etc) | 
| 2386 | 2027 | in case subc mss trec of | 
| 2028 | None => Some(trec) | |
| 2029 | | trec => trec | |
| 2030 | end | |
| 2031 | | _ => | |
| 2032 | let fun appc() = | |
| 4713 | 2033 | (case botc true mss (t,etc) of | 
| 2034 | Some(t1,etc1) => | |
| 2035 | (case botc true mss (u,etc1) of | |
| 2036 | Some(u1,etc2) => Some(t1$u1, etc2) | |
| 2037 | | None => Some(t1$u, etc1)) | |
| 2386 | 2038 | | None => | 
| 4713 | 2039 | (case botc true mss (u,etc) of | 
| 2040 | Some(u1,etc1) => Some(t$u1, etc1) | |
| 2386 | 2041 | | None => None)) | 
| 2042 | val (h,ts) = strip_comb t | |
| 2043 | in case h of | |
| 2044 | Const(a,_) => | |
| 2045 | (case assoc_string(congs,a) of | |
| 2046 | None => appc() | |
| 4116 | 2047 | | Some(cong) => | 
| 2048 | (congc (prover mss,sign_ref,maxidx) cong trec | |
| 2049 | handle Pattern.MATCH => appc() ) ) | |
| 2386 | 2050 | | _ => appc() | 
| 2051 | end) | |
| 2052 | | _ => None) | |
| 0 | 2053 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2054 | and impc args = | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2055 | if mutsimp | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2056 | then let val (prem, conc, mss, etc) = args | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2057 | in snd(mut_impc([], prem, conc, mss, etc)) end | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2058 | else nonmut_impc args | 
| 4713 | 2059 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2060 | and mut_impc (prems, prem, conc, mss, etc) = | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2061 | let val (prem1,etc1) = try_botc mss (prem,etc) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2062 | in mut_impc1(prems, prem1, conc, mss, etc1) end | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2063 | |
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2064 | and mut_impc1(prems, prem1, conc, mss, etc1 as (_,hyps1,_)) = | 
| 4713 | 2065 | let | 
| 4820 
8f6dbbd8d497
Tried to speed up the rewriter by eta-contracting all patterns beforehand and
 nipkow parents: 
4785diff
changeset | 2066 |         fun uncond({thm,lhs,perm}) =
 | 
| 4713 | 2067 | if nprems_of thm = 0 then Some lhs else None | 
| 2068 | ||
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2069 | val (lhss1,mss1) = | 
| 4713 | 2070 | if maxidx_of_term prem1 <> ~1 | 
| 2071 | then (trace_term true "Cannot add premise as rewrite rule because it contains (type) unknowns:" | |
| 2072 | (Sign.deref sign_ref) prem1; | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2073 | ([],mss)) | 
| 4713 | 2074 |           else let val thm = assume (Cterm{sign_ref=sign_ref, t=prem1, 
 | 
| 2075 | T=propT, maxidx= ~1}) | |
| 2076 | val rrules1 = extract_safe_rrules(mss,thm) | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2077 | val lhss1 = mapfilter uncond rrules1 | 
| 4713 | 2078 | val mss1 = foldl insert_rrule (add_prems(mss,[thm]),rrules1) | 
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2079 | in (lhss1, mss1) end | 
| 4713 | 2080 | |
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2081 | fun disch1(conc2,(shyps2,hyps2,ders2)) = | 
| 4713 | 2082 | let val hyps2' = if gen_mem (op aconv) (prem1, hyps1) | 
| 2083 | then hyps2 else hyps2\prem1 | |
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2084 | in (Logic.mk_implies(prem1,conc2),(shyps2,hyps2',ders2)) end | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2085 | |
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2086 | fun rebuild trec2 = | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2087 | let val trec = disch1 trec2 | 
| 4713 | 2088 | in case rewritec (prover,sign_ref,maxidx) mss trec of | 
| 2089 | None => (None,trec) | |
| 2090 |              | Some(Const("==>",_)$prem$conc,etc) =>
 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2091 | mut_impc(prems,prem,conc,mss,etc) | 
| 4713 | 2092 | | Some(trec') => (None,trec') | 
| 2093 | end | |
| 2094 | ||
| 2095 | fun simpconc() = | |
| 2096 | case conc of | |
| 2097 |             Const("==>",_)$s$t =>
 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2098 | (case mut_impc(prems@[prem1],s,t,mss1,etc1) of | 
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2099 | (Some(i,prem),trec2) => | 
| 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2100 | let val trec2' = disch1 trec2 | 
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2101 | in if i=0 then mut_impc1(prems,prem,fst trec2',mss,snd trec2') | 
| 4716 
a291e858061c
Asm_full_simp_tac now reorients asm c = t to t = c.
 nipkow parents: 
4713diff
changeset | 2102 | else (Some(i-1,prem),trec2') | 
| 4713 | 2103 | end | 
| 2104 | | (None,trec) => rebuild(trec)) | |
| 2105 | | _ => rebuild(try_botc mss1 (conc,etc1)) | |
| 2106 | ||
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2107 | in let val sg = Sign.deref sign_ref | 
| 4713 | 2108 | val tsig = #tsig(Sign.rep_sg sg) | 
| 2109 | fun reducible t = | |
| 2110 | exists (fn lhs => Pattern.matches_subterm tsig (lhs,t)) | |
| 2111 | lhss1; | |
| 2112 | in case dropwhile (not o reducible) prems of | |
| 2113 | [] => simpconc() | |
| 4785 | 2114 | | red::rest => (trace_term false "Can now reduce premise:" sg | 
| 4713 | 2115 | red; | 
| 2116 | (Some(length rest,prem1),(conc,etc1))) | |
| 2117 | end | |
| 2118 | end | |
| 0 | 2119 | |
| 4740 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2120 | (* legacy code - only for backwards compatibility *) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2121 | and nonmut_impc(prem, conc, mss, etc as (_,hyps1,_)) = | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2122 | let val (prem1,etc1) = if simprem then try_botc mss (prem,etc) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2123 | else (prem,etc) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2124 | val maxidx1 = maxidx_of_term prem1 | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2125 | val mss1 = | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2126 | if not useprem then mss else | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2127 | if maxidx1 <> ~1 | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2128 | then (trace_term true "Cannot add premise as rewrite rule because it contains (type) unknowns:" | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2129 | (Sign.deref sign_ref) prem1; | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2130 | mss) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2131 |              else let val thm = assume (Cterm{sign_ref=sign_ref, t=prem1, 
 | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2132 | T=propT, maxidx= ~1}) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2133 | in add_safe_simp(add_prems(mss,[thm]), thm) end | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2134 | val (conc2,(shyps2,hyps2,ders2)) = try_botc mss1 (conc,etc1) | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2135 | val hyps2' = if prem1 mem hyps1 then hyps2 else hyps2\prem1 | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2136 | in (Logic.mk_implies(prem1,conc2), (shyps2, hyps2', ders2)) end | 
| 
0136b5bbe9fe
Made mutual simplification of prems a special case.
 nipkow parents: 
4716diff
changeset | 2137 | |
| 1529 | 2138 | in try_botc end; | 
| 0 | 2139 | |
| 2140 | ||
| 2141 | (*** Meta-rewriting: rewrites t to u and returns the theorem t==u ***) | |
| 2509 | 2142 | |
| 2143 | (* | |
| 2144 | Parameters: | |
| 4713 | 2145 | mode = (simplify A, | 
| 2146 | use A in simplifying B, | |
| 2147 | use prems of B (if B is again a meta-impl.) to simplify A) | |
| 2148 | when simplifying A ==> B | |
| 2509 | 2149 | mss: contains equality theorems of the form [|p1,...|] ==> t==u | 
| 2150 | prover: how to solve premises in conditional rewrites and congruences | |
| 0 | 2151 | *) | 
| 2509 | 2152 | |
| 2153 | (* FIXME: check that #bounds(mss) does not "occur" in ct alread *) | |
| 2154 | ||
| 214 
ed6a3e2b1a33
added new parameter to the simplification tactics which indicates if
 nipkow parents: 
209diff
changeset | 2155 | fun rewrite_cterm mode mss prover ct = | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2156 |   let val Cterm {sign_ref, t, T, maxidx} = ct;
 | 
| 4713 | 2157 | val (u,(shyps,hyps,ders)) = bottomc (mode,prover, sign_ref, maxidx) mss | 
| 2158 | (t, (add_term_sorts(t,[]), [], [])); | |
| 0 | 2159 | val prop = Logic.mk_equals(t,u) | 
| 1258 | 2160 | in | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2161 |       Thm{sign_ref = sign_ref, 
 | 
| 2386 | 2162 | der = infer_derivs (Rewrite_cterm ct, ders), | 
| 4116 | 2163 | maxidx = maxidx, | 
| 2386 | 2164 | shyps = shyps, | 
| 2165 | hyps = hyps, | |
| 1529 | 2166 | prop = prop} | 
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2167 | end; | 
| 0 | 2168 | |
| 1539 | 2169 | |
| 2509 | 2170 | |
| 2171 | (*** Oracles ***) | |
| 2172 | ||
| 3812 | 2173 | fun invoke_oracle thy raw_name = | 
| 2174 | let | |
| 2175 |     val {sign = sg, oracles, ...} = rep_theory thy;
 | |
| 2176 | val name = Sign.intern sg Theory.oracleK raw_name; | |
| 2177 | val oracle = | |
| 2178 | (case Symtab.lookup (oracles, name) of | |
| 2179 |         None => raise THM ("Unknown oracle: " ^ name, 0, [])
 | |
| 2180 | | Some (f, _) => f); | |
| 2181 | in | |
| 2182 | fn (sign, exn) => | |
| 2183 | let | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2184 | val sign_ref' = Sign.merge_refs (Sign.self_ref sg, Sign.self_ref sign); | 
| 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2185 | val sign' = Sign.deref sign_ref'; | 
| 3812 | 2186 | val (prop, T, maxidx) = Sign.certify_term sign' (oracle (sign', exn)); | 
| 2187 | in | |
| 2188 | if T <> propT then | |
| 2189 |           raise THM ("Oracle's result must have type prop: " ^ name, 0, [])
 | |
| 2190 | else fix_shyps [] [] | |
| 3967 
edd5ff9371f8
sg_ref: automatic adjustment of thms of draft theories;
 wenzelm parents: 
3895diff
changeset | 2191 |           (Thm {sign_ref = sign_ref', 
 | 
| 4182 | 2192 | der = Join (Oracle (name, sign, exn), []), | 
| 3812 | 2193 | maxidx = maxidx, | 
| 2194 | shyps = [], | |
| 2195 | hyps = [], | |
| 2196 | prop = prop}) | |
| 2197 | end | |
| 2198 | end; | |
| 2199 | ||
| 1539 | 2200 | |
| 0 | 2201 | end; | 
| 1503 | 2202 | |
| 2203 | open Thm; |