| author | huffman | 
| Wed, 12 Oct 2005 03:02:18 +0200 | |
| changeset 17840 | 11bcd77cfa22 | 
| parent 17797 | 9996f3a0ffdf | 
| child 17931 | 881274f283cf | 
| permissions | -rw-r--r-- | 
| 17225 | 1 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | 
| 16179 | 2 | (* Title: Pure/IsaPlanner/term_lib.ML | 
| 3 | ID: $Id$ | |
| 15481 | 4 | Author: Lucas Dixon, University of Edinburgh | 
| 5 | lucasd@dai.ed.ac.uk | |
| 6 | Created: 17 Aug 2002 | |
| 7 | *) | |
| 17225 | 8 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | 
| 15481 | 9 | (* DESCRIPTION: | 
| 10 | ||
| 11 | Additional code to work with terms. | |
| 12 | ||
| 17225 | 13 | *) | 
| 15481 | 14 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | 
| 15 | signature TERM_LIB = | |
| 16 | sig | |
| 17225 | 17 | val fo_term_height : term -> int | 
| 18 | val ho_term_height : term -> int | |
| 15481 | 19 | |
| 20 | (* Matching/unification with exceptions handled *) | |
| 17225 | 21 | val clean_match : theory -> term * term | 
| 22 | -> ((indexname * (sort * typ)) list | |
| 23 | * (indexname * (typ * term)) list) option | |
| 24 | val clean_unify : theory -> int -> term * term | |
| 25 | -> ((indexname * (sort * typ)) list | |
| 26 | * (indexname * (typ * term)) list) Seq.seq | |
| 15481 | 27 | |
| 28 | (* managing variables in terms, can doing conversions *) | |
| 17225 | 29 | val bounds_to_frees : term -> term | 
| 15481 | 30 | val bounds_to_frees_with_vars : | 
| 17225 | 31 | (string * typ) list -> term -> term | 
| 15481 | 32 | |
| 17225 | 33 | val change_bounds_to_frees : term -> term | 
| 34 | val change_frees_to_vars : term -> term | |
| 35 | val change_vars_to_frees : term -> term | |
| 15481 | 36 | val loose_bnds_to_frees : | 
| 17225 | 37 | (string * typ) list -> term -> term | 
| 15481 | 38 | |
| 39 | (* make all variables named uniquely *) | |
| 17225 | 40 | val unique_namify : term -> term | 
| 15481 | 41 | |
| 17225 | 42 | (* breasking up a term and putting it into a normal form | 
| 15481 | 43 | independent of internal term context *) | 
| 17225 | 44 | val cleaned_term_conc : term -> term | 
| 45 | val cleaned_term_parts : term -> term list * term | |
| 46 | val cterm_of_term : term -> cterm | |
| 15481 | 47 | |
| 48 | (* terms of theorems and subgoals *) | |
| 17225 | 49 | val term_of_thm : thm -> term | 
| 50 | val get_prems_of_sg_term : term -> term list | |
| 51 | val triv_thm_from_string : string -> thm | |
| 15481 | 52 | |
| 53 | (* some common term manipulations *) | |
| 17225 | 54 | val try_dest_Goal : term -> term | 
| 55 | val try_dest_Trueprop : term -> term | |
| 15481 | 56 | |
| 17225 | 57 | val bot_left_leaf_of : term -> term | 
| 58 | val bot_left_leaf_noabs_of : term -> term | |
| 15481 | 59 | |
| 60 | (* term containing another term - an embedding check *) | |
| 17225 | 61 | val term_contains : term -> term -> bool | 
| 15481 | 62 | |
| 63 | (* name-convertable checking - like ae-convertable, but allows for | |
| 64 | var's and free's to be mixed - and doesn't used buggy code. :-) *) | |
| 17225 | 65 | val get_name_eq_member : term -> term list -> term option | 
| 66 | val name_eq_member : term -> term list -> bool | |
| 15481 | 67 | val term_name_eq : | 
| 17225 | 68 | term -> | 
| 69 | term -> | |
| 70 | (term * term) list -> | |
| 71 | (term * term) list option | |
| 15481 | 72 | |
| 73 | (* is the typ a function or is it atomic *) | |
| 17225 | 74 | val is_fun_typ : typ -> bool | 
| 75 | val is_atomic_typ : typ -> bool | |
| 15481 | 76 | |
| 77 | (* variable analysis *) | |
| 17225 | 78 | val is_some_kind_of_var : term -> bool | 
| 15481 | 79 | val same_var_check : | 
| 15531 | 80 |        ''a -> ''b -> (''a * ''b) list -> (''a * ''b) list option
 | 
| 17225 | 81 | val has_new_vars : term * term -> bool | 
| 82 | val has_new_typ_vars : term * term -> bool | |
| 15481 | 83 | (* checks to see if the lhs -> rhs is a invalid rewrite rule *) | 
| 17225 | 84 | val is_not_valid_rwrule : theory -> (term * term) -> bool | 
| 15481 | 85 | |
| 86 | (* get the frees in a term that are of atomic type, ie non-functions *) | |
| 17225 | 87 | val atomic_frees_of_term : term -> (string * typ) list | 
| 15481 | 88 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 89 | (* get used names in a theorem to avoid upon instantiation. *) | 
| 17225 | 90 | val usednames_of_term : term -> string list | 
| 91 | val usednames_of_thm : thm -> string list | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 92 | |
| 15481 | 93 | (* Isar term skolemisationm and unsolemisation *) | 
| 94 | (* I think this is written also in IsarRTechn and also in somewhere else *) | |
| 17225 | 95 | (* val skolemise_term : (string,typ) list -> term -> term *) | 
| 96 | val unskolemise_all_term : | |
| 97 | term -> | |
| 98 | (((string * typ) * string) list * term) | |
| 15481 | 99 | |
| 17225 | 100 | (* given a string describing term then a string for its type, returns | 
| 15481 | 101 | read term *) | 
| 17225 | 102 | val readwty : string -> string -> term | 
| 15481 | 103 | |
| 104 | (* pretty stuff *) | |
| 17225 | 105 | val print_term : term -> unit | 
| 15481 | 106 | val pretty_print_sort : string list -> string | 
| 17225 | 107 | val pretty_print_term : term -> string | 
| 108 | val pretty_print_type : typ -> string | |
| 15481 | 109 | val pretty_print_typelist : | 
| 17225 | 110 | typ list -> (typ -> string) -> string | 
| 111 | ||
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 112 | (* for debugging: quickly get a string of a term w.r.t the_context() *) | 
| 17225 | 113 | val string_of_term : term -> string | 
| 15481 | 114 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 115 | (* Pretty printing in a way that allows them to be parsed by ML. | 
| 17225 | 116 | these are perhaps redundent, check the standard basis | 
| 15481 | 117 | lib for generic versions for any datatype? *) | 
| 118 | val writesort : string list -> unit | |
| 17225 | 119 | val writeterm : term -> unit | 
| 120 | val writetype : typ -> unit | |
| 15481 | 121 | end | 
| 122 | ||
| 123 | ||
| 17225 | 124 | structure TermLib : TERM_LIB = | 
| 15481 | 125 | struct | 
| 126 | ||
| 17225 | 127 | (* Two kinds of depth measure for HOAS terms, a first order (flat) and a | 
| 128 | higher order one. | |
| 129 | Note: not stable of eta-contraction: embedding eta-expands term, | |
| 15481 | 130 | thus we assume eta-expanded *) | 
| 17225 | 131 | fun fo_term_height (a $ b) = | 
| 17797 | 132 | Int.max (1 + fo_term_height b, (fo_term_height a)) | 
| 15481 | 133 | | fo_term_height (Abs(_,_,t)) = fo_term_height t | 
| 134 | | fo_term_height _ = 0; | |
| 17225 | 135 | |
| 136 | fun ho_term_height (a $ b) = | |
| 17797 | 137 | 1 + (Int.max (ho_term_height b, ho_term_height a)) | 
| 15481 | 138 | | ho_term_height (Abs(_,_,t)) = ho_term_height t | 
| 139 | | ho_term_height _ = 0; | |
| 140 | ||
| 15928 
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
 dixon parents: 
15915diff
changeset | 141 | |
| 15481 | 142 | (* Higher order matching with exception handled *) | 
| 143 | (* returns optional instantiation *) | |
| 17225 | 144 | fun clean_match thy (a as (pat, t)) = | 
| 17203 | 145 | let val (tyenv, tenv) = Pattern.match thy a | 
| 15915 
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
 dixon parents: 
15814diff
changeset | 146 | in SOME (Vartab.dest tyenv, Vartab.dest tenv) | 
| 15798 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15570diff
changeset | 147 | end handle Pattern.MATCH => NONE; | 
| 15481 | 148 | (* Higher order unification with exception handled, return the instantiations *) | 
| 149 | (* given Signature, max var index, pat, tgt *) | |
| 150 | (* returns Seq of instantiations *) | |
| 17225 | 151 | fun clean_unify sgn ix (a as (pat, tgt)) = | 
| 152 | let | |
| 153 | (* type info will be re-derived, maybe this can be cached | |
| 15481 | 154 | for efficiency? *) | 
| 155 | val pat_ty = Term.type_of pat; | |
| 156 | val tgt_ty = Term.type_of tgt; | |
| 17225 | 157 | (* is it OK to ignore the type instantiation info? | 
| 15481 | 158 | or should I be using it? *) | 
| 17225 | 159 | val typs_unify = | 
| 16934 | 160 | SOME (Sign.typ_unify sgn (pat_ty, tgt_ty) (Term.Vartab.empty, ix)) | 
| 161 | handle Type.TUNIFY => NONE; | |
| 15481 | 162 | in | 
| 163 | case typs_unify of | |
| 15531 | 164 | SOME (typinsttab, ix2) => | 
| 17225 | 165 | let | 
| 166 | (* is it right to throw away the flexes? | |
| 15481 | 167 | or should I be using them somehow? *) | 
| 17225 | 168 | fun mk_insts env = | 
| 15915 
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
 dixon parents: 
15814diff
changeset | 169 | (Vartab.dest (Envir.type_env env), | 
| 
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
 dixon parents: 
15814diff
changeset | 170 | Envir.alist_of env); | 
| 17225 | 171 |           val initenv = Envir.Envir {asol = Vartab.empty,
 | 
| 15481 | 172 | iTs = typinsttab, maxidx = ix2}; | 
| 15928 
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
 dixon parents: 
15915diff
changeset | 173 | val useq = (Unify.smash_unifiers (sgn,initenv,[a])) | 
| 
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
 dixon parents: 
15915diff
changeset | 174 | handle UnequalLengths => Seq.empty | 
| 
66b165ee016c
lucas - made clean unify smash unifiers so that when we get flex-flex constraints subst does not barf. Also added fix_vars_upto_idx to IsaND.
 dixon parents: 
15915diff
changeset | 175 | | Term.TERM _ => Seq.empty; | 
| 17225 | 176 | fun clean_unify' useq () = | 
| 177 | (case (Seq.pull useq) of | |
| 15531 | 178 | NONE => NONE | 
| 179 | | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) | |
| 15570 | 180 | handle UnequalLengths => NONE | 
| 15531 | 181 | | Term.TERM _ => NONE; | 
| 15481 | 182 | in | 
| 183 | (Seq.make (clean_unify' useq)) | |
| 184 | end | |
| 15531 | 185 | | NONE => Seq.empty | 
| 15481 | 186 | end; | 
| 187 | ||
| 17225 | 188 | fun asm_mk t = assume (cterm_of (the_context()) t); | 
| 189 | fun asm_read s = (Thm.assume (read_cterm (the_context()) (s,propT))); | |
| 15481 | 190 | |
| 191 | ||
| 192 | (* more pretty printing code for Isabelle terms etc *) | |
| 193 | ||
| 194 | ||
| 17225 | 195 | (* pretty_print_typelist l f = print a typelist. | 
| 15481 | 196 | l = list of types to print : typ list | 
| 197 | f = function used to print a single type : typ -> string | |
| 198 | *) | |
| 199 | fun pretty_print_typelist [] f = "" | |
| 200 | | pretty_print_typelist [(h: typ)] (f : typ -> string) = (f h) | |
| 17225 | 201 | | pretty_print_typelist ((h: typ) :: t) (f : typ -> string) = | 
| 15481 | 202 | (f h) ^ ", " ^ (pretty_print_typelist t f); | 
| 203 | ||
| 204 | ||
| 205 | ||
| 17225 | 206 | (* pretty_print_sort s = print a sort | 
| 15481 | 207 | s = sort to print : string list | 
| 208 | *) | |
| 209 | fun pretty_print_sort [] = "" | |
| 210 | | pretty_print_sort ([h]) = "\"" ^ h ^ "\"" | |
| 211 | | pretty_print_sort (h :: t) = "\"" ^ h ^ "\"," ^ (pretty_print_sort t); | |
| 212 | ||
| 213 | ||
| 214 | (* pretty_print_type t = print a type | |
| 215 | t = type to print : type | |
| 216 | *) | |
| 17225 | 217 | fun pretty_print_type (Type (n, l)) = | 
| 15481 | 218 | "Type(\"" ^ n ^ "\", [" ^ (pretty_print_typelist l pretty_print_type) ^ "])" | 
| 17225 | 219 | | pretty_print_type (TFree (n, s)) = | 
| 15481 | 220 | "TFree(\"" ^ n ^ "\", [" ^ (pretty_print_sort s) ^ "])" | 
| 17225 | 221 | | pretty_print_type (TVar ((n, i), s)) = | 
| 15481 | 222 | "TVar( (\"" ^ n ^ "\", " ^ (string_of_int i) ^ "), [" ^ (pretty_print_sort s) ^ "])"; | 
| 223 | ||
| 224 | ||
| 225 | (* pretty_print_term t = print a term prints types and sorts too. | |
| 226 | t = term to print : term | |
| 227 | *) | |
| 17225 | 228 | fun pretty_print_term (Const (s, t)) = | 
| 15481 | 229 | "Const(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" | 
| 17225 | 230 | | pretty_print_term (Free (s, t)) = | 
| 15481 | 231 | "Free(\"" ^ s ^ "\", " ^ (pretty_print_type t) ^ ")" | 
| 17225 | 232 | | pretty_print_term (Var ((n, i), t)) = | 
| 15481 | 233 | "Var( (\"" ^ n ^ "\"," ^ (string_of_int i) ^ "), " ^ (pretty_print_type t) ^ ")" | 
| 17225 | 234 | | pretty_print_term (Bound i) = | 
| 15481 | 235 |       "Bound(" ^ (string_of_int i) ^ ")"
 | 
| 17225 | 236 | | pretty_print_term (Abs (s, t, r)) = | 
| 15481 | 237 | "Abs(\"" ^ s ^ "\"," ^ (pretty_print_type t) ^ ", \n " ^ (pretty_print_term r) ^ ")" | 
| 17225 | 238 | | pretty_print_term (op $ (t1, t2)) = | 
| 15481 | 239 |       "(" ^ (pretty_print_term t1) ^ ") $\n (" ^ (pretty_print_term t2) ^ ")";
 | 
| 240 | ||
| 241 | (* Write the term out nicly instead of just creating a string for it *) | |
| 242 | fun writeterm t = writeln (pretty_print_term t); | |
| 243 | fun writetype t = writeln (pretty_print_type t); | |
| 244 | fun writesort s = writeln (pretty_print_sort s); | |
| 245 | ||
| 17225 | 246 | fun cterm_of_term (t : term) = Thm.cterm_of (the_context()) t; | 
| 15481 | 247 | fun term_of_thm t = (Thm.prop_of t); | 
| 248 | ||
| 17225 | 249 | fun string_of_term t = Sign.string_of_term (the_context()) t; | 
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 250 | fun print_term t = writeln (string_of_term t); | 
| 15481 | 251 | |
| 252 | (* create a trivial HOL thm from anything... *) | |
| 17225 | 253 | fun triv_thm_from_string s = Thm.trivial (cterm_of (the_context()) (read s)); | 
| 15481 | 254 | |
| 255 | (* Checks if vars could be the same - alpha convertable | |
| 256 | w.r.t. previous vars, a and b are assumed to be vars, | |
| 257 | free vars, but not bound vars, | |
| 258 | Note frees and vars must all have unique names. *) | |
| 259 | fun same_var_check a b L = | |
| 17225 | 260 | let | 
| 15531 | 261 | fun bterm_from t [] = NONE | 
| 17225 | 262 | | bterm_from t ((a,b)::m) = | 
| 15531 | 263 | if t = a then SOME b else bterm_from t m | 
| 15481 | 264 | |
| 265 | val bl_opt = bterm_from a L | |
| 266 | in | |
| 267 | case bterm_from a L of | |
| 15531 | 268 | NONE => SOME ((a,b)::L) | 
| 269 | | SOME b2 => if b2 = b then SOME L else NONE | |
| 15481 | 270 | end; | 
| 271 | ||
| 272 | (* FIXME: make more efficient, only require a single new var! *) | |
| 273 | (* check if the new term has any meta variables not in the old term *) | |
| 17225 | 274 | fun has_new_vars (old, new) = | 
| 275 | (case IsaPLib.lrem (op =) (Term.term_vars old) (Term.term_vars new) of | |
| 15481 | 276 | [] => false | 
| 277 | | (_::_) => true); | |
| 278 | (* check if the new term has any meta variables not in the old term *) | |
| 17225 | 279 | fun has_new_typ_vars (old, new) = | 
| 280 | (case IsaPLib.lrem (op =) (Term.term_tvars old) (Term.term_tvars new) of | |
| 15481 | 281 | [] => false | 
| 282 | | (_::_) => true); | |
| 283 | ||
| 15570 | 284 | (* This version avoids name conflicts that might be introduced by | 
| 15481 | 285 | unskolemisation, and returns a list of (string * Term.typ) * string, | 
| 286 | where the outer string is the original name, and the inner string is | |
| 287 | the new name, and the type is the type of the free variable that was | |
| 288 | renamed. *) | |
| 289 | local | |
| 17225 | 290 | fun myadd (n,ty) (L as (renames, names)) = | 
| 291 | let val n' = Syntax.dest_skolem n in | |
| 292 | case (Library.find_first (fn (_,n2) => (n2 = n)) renames) of | |
| 293 | NONE => | |
| 294 | let val renamedn = Term.variant names n' in | |
| 295 | (renamedn, (((renamedn, ty), n) :: renames, | |
| 15481 | 296 | renamedn :: names)) end | 
| 15531 | 297 | | (SOME ((renamedn, _), _)) => (renamedn, L) | 
| 15481 | 298 | end | 
| 15570 | 299 | handle Fail _ => (n, L); | 
| 15481 | 300 | |
| 17225 | 301 | fun unsk (L,f) (t1 $ t2) = | 
| 302 | let val (L', t1') = unsk (L, I) t1 | |
| 15481 | 303 | in unsk (L', (fn x => f (t1' $ x))) t2 end | 
| 17225 | 304 | | unsk (L,f) (Abs(n,ty,t)) = | 
| 15481 | 305 | unsk (L, (fn x => Abs(n,ty,x))) t | 
| 17225 | 306 | | unsk (L, f) (Free (n,ty)) = | 
| 15481 | 307 | let val (renamed_n, L') = myadd (n ,ty) L | 
| 308 | in (L', f (Free (renamed_n, ty))) end | |
| 309 | | unsk (L, f) l = (L, f l); | |
| 310 | in | |
| 17225 | 311 | fun unskolemise_all_term t = | 
| 312 | let | |
| 313 | val names = Term.add_term_names (t,[]) | |
| 15481 | 314 | val ((renames,names'),t') = unsk (([], names),I) t | 
| 315 | in (renames,t') end; | |
| 316 | end | |
| 317 | ||
| 318 | (* true if the type t is a function *) | |
| 17225 | 319 | fun is_fun_typ (Type(s, l)) = | 
| 15481 | 320 | if s = "fun" then true else false | 
| 321 | | is_fun_typ _ = false; | |
| 322 | ||
| 323 | val is_atomic_typ = not o is_fun_typ; | |
| 324 | ||
| 325 | (* get the frees in a term that are of atomic type, ie non-functions *) | |
| 326 | val atomic_frees_of_term = | |
| 17225 | 327 | List.filter (is_atomic_typ o snd) | 
| 328 | o map Term.dest_Free | |
| 15481 | 329 | o Term.term_frees; | 
| 330 | ||
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 331 | fun usednames_of_term t = Term.add_term_names (t,[]); | 
| 17225 | 332 | fun usednames_of_thm th = | 
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 333 | let val rep = Thm.rep_thm th | 
| 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 334 | val hyps = #hyps rep | 
| 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 335 | val (tpairl,tpairr) = Library.split_list (#tpairs rep) | 
| 17225 | 336 | val prop = #prop rep | 
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 337 | in | 
| 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 338 | List.foldr Term.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps))) | 
| 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 339 | end; | 
| 15481 | 340 | |
| 17225 | 341 | (* read in a string and a top-level type and this will give back a term *) | 
| 342 | fun readwty tstr tystr = | |
| 343 | let val thy = the_context() | |
| 344 | in Sign.simple_read_term thy (Sign.read_typ (thy, K NONE) tystr) tstr end; | |
| 15481 | 345 | |
| 346 | (* first term is equal to the second in some name convertable | |
| 347 | way... Note: This differs from the aeconv in the Term.ML file of | |
| 17225 | 348 | Isabelle in that it allows a var to be alpha-equiv to a free var. | 
| 349 | ||
| 15481 | 350 | Also, the isabelle term.ML version of aeconv uses a | 
| 351 | function that it claims doesn't work! *) | |
| 17225 | 352 | fun term_name_eq (Abs(_,ty1,t1)) (Abs(_,ty2,t2)) l = | 
| 15481 | 353 | if ty1 = ty2 then term_name_eq t1 t2 l | 
| 15531 | 354 | else NONE | 
| 15481 | 355 | | term_name_eq (ah $ at) (bh $ bt) l = | 
| 17225 | 356 | let | 
| 15481 | 357 | val lopt = (term_name_eq ah bh l) | 
| 358 | in | |
| 17225 | 359 | if isSome lopt then | 
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 360 | term_name_eq at bt (valOf lopt) | 
| 15481 | 361 | else | 
| 15531 | 362 | NONE | 
| 15481 | 363 | end | 
| 17225 | 364 | | term_name_eq (Const(a,T)) (Const(b,U)) l = | 
| 15531 | 365 | if a=b andalso T=U then SOME l | 
| 366 | else NONE | |
| 17225 | 367 | | term_name_eq (a as Free(s1,ty1)) (b as Free(s2,ty2)) l = | 
| 15481 | 368 | same_var_check a b l | 
| 17225 | 369 | | term_name_eq (a as Free(s1,ty1)) (b as Var(n2,ty2)) l = | 
| 15481 | 370 | same_var_check a b l | 
| 17225 | 371 | | term_name_eq (a as Var(n1,ty1)) (b as Free(s2,ty2)) l = | 
| 15481 | 372 | same_var_check a b l | 
| 17225 | 373 | | term_name_eq (a as Var(n1,ty1)) (b as Var(n2,ty2)) l = | 
| 15481 | 374 | same_var_check a b l | 
| 17225 | 375 | | term_name_eq (Bound i) (Bound j) l = | 
| 15531 | 376 | if i = j then SOME l else NONE | 
| 377 |   | term_name_eq a b l = ((*writeln ("unchecked case:\n" ^ "a:\n" ^ (pretty_print_term a) ^ "\nb:\n" ^ (pretty_print_term b));*) NONE);
 | |
| 15481 | 378 | |
| 379 | (* checks to see if the term in a name-equivalent member of the list of terms. *) | |
| 15531 | 380 | fun get_name_eq_member a [] = NONE | 
| 17225 | 381 | | get_name_eq_member a (h :: t) = | 
| 382 | if isSome (term_name_eq a h []) then SOME h | |
| 15481 | 383 | else get_name_eq_member a t; | 
| 384 | ||
| 385 | fun name_eq_member a [] = false | |
| 17225 | 386 | | name_eq_member a (h :: t) = | 
| 387 | if isSome (term_name_eq a h []) then true | |
| 15481 | 388 | else name_eq_member a t; | 
| 389 | ||
| 390 | (* true if term is a variable *) | |
| 391 | fun is_some_kind_of_var (Free(s, ty)) = true | |
| 392 | | is_some_kind_of_var (Var(i, ty)) = true | |
| 393 | | is_some_kind_of_var (Bound(i)) = true | |
| 394 | | is_some_kind_of_var _ = false; | |
| 395 | ||
| 396 | (* checks to see if the lhs -> rhs is a invalid rewrite rule *) | |
| 397 | (* FIXME: we should really check that there is a subterm on the lhs | |
| 398 | which embeds into the rhs, this would be much closer to the normal | |
| 399 | notion of valid wave rule - ie there exists at least one case where it | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 400 | is a valid wave rule... *) | 
| 17225 | 401 | fun is_not_valid_rwrule thy (lhs, rhs) = | 
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 402 | Term.is_Var (Term.head_of lhs) (* if lhs is essentially just a var *) | 
| 17225 | 403 | orelse has_new_vars (lhs,rhs) | 
| 404 | orelse has_new_typ_vars (lhs,rhs) | |
| 17203 | 405 | orelse Pattern.matches_subterm thy (lhs, rhs); | 
| 15481 | 406 | |
| 407 | ||
| 408 | (* first term contains the second in some name convertable way... *) | |
| 409 | (* note: this is equivalent to an alpha-convertable emedding *) | |
| 410 | (* takes as args: term containing a, term contained b, | |
| 17225 | 411 | (bound vars of a, bound vars of b), | 
| 412 | current alpha-conversion-pairs, | |
| 15481 | 413 | returns: bool:can convert, new alpha-conversion table *) | 
| 414 | (* in bellow: we *don't* use: a loose notion that only requires | |
| 415 | variables to match variables, and doesn't worry about the actual | |
| 416 | pattern in the variables. *) | |
| 17225 | 417 | fun term_contains1 (Bs, FVs) (Abs(s,ty,t)) (Abs(s2,ty2,t2)) = | 
| 418 | if ty = ty2 then | |
| 15531 | 419 | term_contains1 ((SOME(s,s2,ty)::Bs), FVs) t t2 | 
| 15481 | 420 | else [] | 
| 421 | ||
| 422 | | term_contains1 T t1 (Abs(s2,ty2,t2)) = [] | |
| 423 | ||
| 17225 | 424 | | term_contains1 (Bs, FVs) (Abs(s,ty,t)) t2 = | 
| 15531 | 425 | term_contains1 (NONE::Bs, FVs) t t2 | 
| 15481 | 426 | |
| 427 | | term_contains1 T (ah $ at) (bh $ bt) = | |
| 17225 | 428 | (term_contains1 T ah (bh $ bt)) @ | 
| 15481 | 429 | (term_contains1 T at (bh $ bt)) @ | 
| 17225 | 430 | (List.concat (map (fn inT => (term_contains1 inT at bt)) | 
| 15481 | 431 | (term_contains1 T ah bh))) | 
| 432 | ||
| 433 | | term_contains1 T a (bh $ bt) = [] | |
| 434 | ||
| 435 | | term_contains1 T (ah $ at) b = | |
| 436 | (term_contains1 T ah b) @ (term_contains1 T at b) | |
| 437 | ||
| 17225 | 438 | | term_contains1 T a b = | 
| 15481 | 439 | (* simple list table lookup to check if a named variable has been | 
| 440 | mapped to a variable, if not adds the mapping and return some new | |
| 441 | list mapping, if it is then it checks that the pair are mapped to | |
| 442 | each other, if so returns the current mapping list, else none. *) | |
| 17225 | 443 | let | 
| 15531 | 444 | fun bterm_from t [] = NONE | 
| 17225 | 445 | | bterm_from t ((a,b)::m) = | 
| 15531 | 446 | if t = a then SOME b else bterm_from t m | 
| 15481 | 447 | |
| 448 | (* check to see if, w.r.t. the variable mapping, two terms are leaf | |
| 449 | terms and are mapped to each other. Note constants are only mapped | |
| 17225 | 450 | to the same constant. *) | 
| 15481 | 451 | fun same_leaf_check (T as (Bs,FVs)) (Bound i) (Bound j) = | 
| 17225 | 452 | let | 
| 15481 | 453 | fun aux_chk (i1,i2) [] = false | 
| 15531 | 454 | | aux_chk (0,0) ((SOME _) :: bnds) = true | 
| 455 | | aux_chk (i1,0) (NONE :: bnds) = false | |
| 456 | | aux_chk (i1,i2) ((SOME _) :: bnds) = | |
| 15481 | 457 | aux_chk (i1 - 1,i2 - 1) bnds | 
| 15531 | 458 | | aux_chk (i1,i2) (NONE :: bnds) = | 
| 17225 | 459 | aux_chk (i1,i2 - 1) bnds | 
| 15481 | 460 | in | 
| 461 | if (aux_chk (i,j) Bs) then [T] | |
| 462 | else [] | |
| 463 | end | |
| 17225 | 464 | | same_leaf_check (T as (Bs,(Fs,Vs))) | 
| 465 | (a as (Free (an,aty))) | |
| 15481 | 466 | (b as (Free (bn,bty))) = | 
| 17225 | 467 | (case bterm_from an Fs of | 
| 15531 | 468 | SOME b2n => if bn = b2n then [T] | 
| 15481 | 469 | else [] (* conflict of var name *) | 
| 15531 | 470 | | NONE => [(Bs,((an,bn)::Fs,Vs))]) | 
| 17225 | 471 | | same_leaf_check (T as (Bs,(Fs,Vs))) | 
| 472 | (a as (Var (an,aty))) | |
| 15481 | 473 | (b as (Var (bn,bty))) = | 
| 17225 | 474 | (case bterm_from an Vs of | 
| 15531 | 475 | SOME b2n => if bn = b2n then [T] | 
| 15481 | 476 | else [] (* conflict of var name *) | 
| 15531 | 477 | | NONE => [(Bs,(Fs,(an,bn)::Vs))]) | 
| 15481 | 478 | | same_leaf_check T (a as (Const _)) (b as (Const _)) = | 
| 479 | if a = b then [T] else [] | |
| 480 | | same_leaf_check T _ _ = [] | |
| 481 | ||
| 482 | in | |
| 483 | same_leaf_check T a b | |
| 484 | end; | |
| 485 | ||
| 486 | (* wrapper for term_contains1: checks if the term "a" contains in | |
| 487 | some embedded way, (w.r.t. name -convertable) the term "b" *) | |
| 17225 | 488 | fun term_contains a b = | 
| 15481 | 489 | case term_contains1 ([],([],[])) a b of | 
| 490 | (_ :: _) => true | |
| 491 | | [] => false; | |
| 492 | ||
| 493 | (* change all bound variables to see ones with appropriate name and | |
| 494 | type *) | |
| 495 | (* naming convention is OK as we use 'variant' from term.ML *) | |
| 496 | (* Note "bounds_to_frees" defined below, its better and quicker, but | |
| 497 | keeps the quantifiers handing about, and changes all bounds, not | |
| 498 | just universally quantified ones. *) | |
| 17225 | 499 | fun change_bounds_to_frees t = | 
| 500 | let | |
| 15481 | 501 | val vars = strip_all_vars t | 
| 502 | val frees_names = map (fn Free(s,n) => s | _ => "") (term_frees t) | |
| 503 | val body = strip_all_body t | |
| 504 | ||
| 505 | fun bnds_to_frees [] _ acc = acc | |
| 17225 | 506 | | bnds_to_frees ((name,ty)::more) names acc = | 
| 507 | let | |
| 15481 | 508 | val new_name = variant names name | 
| 509 | in | |
| 510 | bnds_to_frees more (new_name::names) (Free(new_name,ty)::acc) | |
| 511 | end | |
| 512 | in | |
| 513 | (subst_bounds ((bnds_to_frees vars frees_names []), body)) | |
| 17225 | 514 | end; | 
| 15481 | 515 | |
| 516 | (* a runtime-quick function which makes sure that every variable has a | |
| 517 | unique name *) | |
| 17225 | 518 | fun unique_namify_aux (ntab,(Abs(s,ty,t))) = | 
| 17412 | 519 | (case Symtab.lookup ntab s of | 
| 17225 | 520 | NONE => | 
| 521 | let | |
| 17412 | 522 | val (ntab2,t2) = unique_namify_aux (Symtab.update (s, s) ntab, t) | 
| 15481 | 523 | in | 
| 17225 | 524 | (ntab2,Abs(s,ty,t2)) | 
| 15481 | 525 | end | 
| 17225 | 526 | | SOME s2 => | 
| 527 | let | |
| 528 | val s_new = (Term.variant (Symtab.keys ntab) s) | |
| 529 | val (ntab2,t2) = | |
| 17412 | 530 | unique_namify_aux (Symtab.update (s_new, s_new) ntab, t) | 
| 15481 | 531 | in | 
| 17225 | 532 | (ntab2,Abs(s_new,ty,t2)) | 
| 15481 | 533 | end) | 
| 17225 | 534 | | unique_namify_aux (ntab,(a $ b)) = | 
| 535 | let | |
| 15481 | 536 | val (ntab1,t1) = unique_namify_aux (ntab,a) | 
| 17225 | 537 | val (ntab2,t2) = unique_namify_aux (ntab1,b) | 
| 15481 | 538 | in | 
| 539 | (ntab2, t1$t2) | |
| 540 | end | |
| 541 | | unique_namify_aux (nt as (ntab,Const x)) = nt | |
| 17225 | 542 | | unique_namify_aux (nt as (ntab,f as Free (s,ty))) = | 
| 17412 | 543 | (case Symtab.lookup ntab s of | 
| 544 | NONE => (Symtab.update (s, s) ntab, f) | |
| 15531 | 545 | | SOME _ => nt) | 
| 17225 | 546 | | unique_namify_aux (nt as (ntab,v as Var ((s,i),ty))) = | 
| 17412 | 547 | (case Symtab.lookup ntab s of | 
| 548 | NONE => (Symtab.update (s, s) ntab, v) | |
| 15531 | 549 | | SOME _ => nt) | 
| 15481 | 550 | | unique_namify_aux (nt as (ntab, Bound i)) = nt; | 
| 551 | ||
| 17225 | 552 | fun unique_namify t = | 
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 553 | #2 (unique_namify_aux (Symtab.empty, t)); | 
| 15481 | 554 | |
| 555 | (* a runtime-quick function which makes sure that every variable has a | |
| 556 | unique name and also changes bound variables to free variables, used | |
| 557 | for embedding checks, Note that this is a pretty naughty term | |
| 558 | manipulation, which doesn't have necessary relation to the original | |
| 559 | sematics of the term. This is really a trick for our embedding code. *) | |
| 560 | ||
| 17225 | 561 | fun bounds_to_frees_aux T (ntab,(Abs(s,ty,t))) = | 
| 17412 | 562 | (case Symtab.lookup ntab s of | 
| 17225 | 563 | NONE => | 
| 564 | let | |
| 565 | val (ntab2,t2) = | |
| 17412 | 566 | bounds_to_frees_aux ((s,ty)::T) (Symtab.update (s, s) ntab, t) | 
| 15481 | 567 | in | 
| 568 | (ntab2,Abs(s,ty,t2)) | |
| 569 | end | |
| 17225 | 570 | | SOME s2 => | 
| 571 | let | |
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 572 | val s_new = (Term.variant (Symtab.keys ntab) s) | 
| 17225 | 573 | val (ntab2,t2) = | 
| 574 | bounds_to_frees_aux ((s_new,ty)::T) | |
| 17412 | 575 | (Symtab.update (s_new, s_new) ntab, t) | 
| 15481 | 576 | in | 
| 577 | (ntab2,Abs(s_new,ty,t2)) | |
| 578 | end) | |
| 17225 | 579 | | bounds_to_frees_aux T (ntab,(a $ b)) = | 
| 580 | let | |
| 15481 | 581 | val (ntab1,t1) = bounds_to_frees_aux T (ntab,a) | 
| 17225 | 582 | val (ntab2,t2) = bounds_to_frees_aux T (ntab1,b) | 
| 15481 | 583 | in | 
| 584 | (ntab2, t1$t2) | |
| 585 | end | |
| 586 | | bounds_to_frees_aux T (nt as (ntab,Const x)) = nt | |
| 17225 | 587 | | bounds_to_frees_aux T (nt as (ntab,f as Free (s,ty))) = | 
| 17412 | 588 | (case Symtab.lookup ntab s of | 
| 589 | NONE => (Symtab.update (s, s) ntab, f) | |
| 15531 | 590 | | SOME _ => nt) | 
| 17225 | 591 | | bounds_to_frees_aux T (nt as (ntab,v as Var ((s,i),ty))) = | 
| 17412 | 592 | (case Symtab.lookup ntab s of | 
| 593 | NONE => (Symtab.update (s, s) ntab, v) | |
| 15531 | 594 | | SOME _ => nt) | 
| 17225 | 595 | | bounds_to_frees_aux T (nt as (ntab, Bound i)) = | 
| 596 | let | |
| 597 | val (s,ty) = List.nth (T,i) | |
| 15481 | 598 | in | 
| 599 | (ntab, Free (s,ty)) | |
| 600 | end; | |
| 601 | ||
| 602 | ||
| 17225 | 603 | fun bounds_to_frees t = | 
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 604 | #2 (bounds_to_frees_aux [] (Symtab.empty,t)); | 
| 15481 | 605 | |
| 17225 | 606 | fun bounds_to_frees_with_vars vars t = | 
| 16857 
6389511d4609
lucas - slightly cleaned up. Removed redudent copy of Symtab structure.
 dixon parents: 
16179diff
changeset | 607 | #2 (bounds_to_frees_aux vars (Symtab.empty,t)); | 
| 15481 | 608 | |
| 609 | ||
| 610 | ||
| 611 | (* loose bounds to frees *) | |
| 17225 | 612 | fun loose_bnds_to_frees_aux (bnds,vars) (Abs(s,ty,t)) = | 
| 15481 | 613 | Abs(s,ty,loose_bnds_to_frees_aux (bnds + 1,vars) t) | 
| 17225 | 614 | | loose_bnds_to_frees_aux (bnds,vars) (a $ b) = | 
| 615 | (loose_bnds_to_frees_aux (bnds,vars) a) | |
| 15481 | 616 | $ (loose_bnds_to_frees_aux (bnds,vars) b) | 
| 17225 | 617 | | loose_bnds_to_frees_aux (bnds,vars) (t as (Bound i)) = | 
| 15570 | 618 | if (bnds <= i) then Free (List.nth (vars,i - bnds)) | 
| 15481 | 619 | else t | 
| 620 | | loose_bnds_to_frees_aux (bnds,vars) t = t; | |
| 621 | ||
| 622 | ||
| 17225 | 623 | fun loose_bnds_to_frees vars t = | 
| 15481 | 624 | loose_bnds_to_frees_aux (0,vars) t; | 
| 625 | ||
| 626 | ||
| 627 |   fun try_dest_Goal (Const("Goal", _) $ T) = T
 | |
| 628 | | try_dest_Goal T = T; | |
| 629 | ||
| 630 |   fun try_dest_Trueprop (Const("Trueprop", _) $ T) = T
 | |
| 17225 | 631 | | try_dest_Trueprop T = T; | 
| 15481 | 632 | |
| 633 | fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l | |
| 634 | | bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t | |
| 635 | | bot_left_leaf_of x = x; | |
| 636 | ||
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 637 | fun bot_left_leaf_noabs_of (l $ r) = bot_left_leaf_noabs_of l | 
| 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15798diff
changeset | 638 | | bot_left_leaf_noabs_of x = x; | 
| 15481 | 639 | |
| 640 | (* cleaned up normal form version of the subgoal terms conclusion *) | |
| 17225 | 641 | fun cleaned_term_conc t = | 
| 15481 | 642 | let | 
| 643 | val concl = Logic.strip_imp_concl (change_bounds_to_frees t) | |
| 17225 | 644 | in | 
| 15481 | 645 | (try_dest_Trueprop (try_dest_Goal concl)) | 
| 646 | end; | |
| 647 | ||
| 17225 | 648 | (* fun get_prems_of_sg_term t = | 
| 15481 | 649 | map opt_dest_Trueprop (Logic.strip_imp_prems t); *) | 
| 650 | ||
| 17225 | 651 | fun get_prems_of_sg_term t = | 
| 15481 | 652 | map try_dest_Trueprop (Logic.strip_assums_hyp t); | 
| 653 | ||
| 654 | ||
| 655 | (* drop premices, clean bound var stuff, and make a trueprop... *) | |
| 17225 | 656 | fun cleaned_term_parts t = | 
| 657 | let | |
| 15481 | 658 | val t2 = (change_bounds_to_frees t) | 
| 659 | val concl = Logic.strip_imp_concl t2 | |
| 660 | val prems = map try_dest_Trueprop (Logic.strip_imp_prems t2) | |
| 17225 | 661 | in | 
| 15481 | 662 | (prems, (try_dest_Trueprop (try_dest_Goal concl))) | 
| 663 | end; | |
| 664 | ||
| 665 | (* change free variables to vars and visa versa *) | |
| 666 | (* *** check naming is OK, can we just use the vasr old name? *) | |
| 667 | (* *** Check: Logic.varify and Logic.unvarify *) | |
| 17225 | 668 | fun change_vars_to_frees (a$b) = | 
| 15481 | 669 | (change_vars_to_frees a) $ (change_vars_to_frees b) | 
| 17225 | 670 | | change_vars_to_frees (Abs(s,ty,t)) = | 
| 15481 | 671 | (Abs(s,Type.varifyT ty,change_vars_to_frees t)) | 
| 672 | | change_vars_to_frees (Var((s,i),ty)) = (Free(s,Type.unvarifyT ty)) | |
| 673 | | change_vars_to_frees l = l; | |
| 674 | ||
| 17225 | 675 | fun change_frees_to_vars (a$b) = | 
| 15481 | 676 | (change_frees_to_vars a) $ (change_frees_to_vars b) | 
| 17225 | 677 | | change_frees_to_vars (Abs(s,ty,t)) = | 
| 15481 | 678 | (Abs(s,Type.varifyT ty,change_frees_to_vars t)) | 
| 679 | | change_frees_to_vars (Free(s,ty)) = (Var((s,0),Type.varifyT ty)) | |
| 680 | | change_frees_to_vars l = l; | |
| 681 | ||
| 682 | ||
| 16934 | 683 | end; |