| 55596 |      1 | (*  Title:      HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML
 | 
|  |      2 |     Author:     Nik Sultana, Cambridge University Computer Laboratory
 | 
|  |      3 | Collection of general functions used in the reconstruction module.
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | signature TPTP_RECONSTRUCT_LIBRARY =
 | 
|  |      7 | sig
 | 
|  |      8 |   exception BREAK_LIST
 | 
|  |      9 |   val break_list : 'a list -> 'a * 'a list
 | 
|  |     10 |   val break_seq : 'a Seq.seq -> 'a * 'a Seq.seq
 | 
|  |     11 |   exception MULTI_ELEMENT_LIST
 | 
|  |     12 |   val cascaded_filter_single : bool -> ('a list -> 'a list) list -> 'a list -> 'a option
 | 
|  |     13 |   val concat_between : 'a list list -> ('a option * 'a option) -> 'a list
 | 
|  |     14 |   exception DIFF_TYPE of typ * typ
 | 
|  |     15 |   exception DIFF of term * term
 | 
|  |     16 |   val diff :
 | 
|  |     17 |      theory ->
 | 
|  |     18 |      term * term -> (term * term) list * (typ * typ) list
 | 
|  |     19 |   exception DISPLACE_KV
 | 
|  |     20 |   val displace_kv : ''a -> (''a * 'b) list -> (''a * 'b) list
 | 
|  |     21 |   val enumerate : int -> 'a list -> (int * 'a) list
 | 
|  |     22 |   val fold_options : 'a option list -> 'a list
 | 
|  |     23 |   val find_and_remove : ('a -> bool) -> 'a list -> 'a * 'a list
 | 
|  |     24 |   val lift_option : ('a -> 'b) -> 'a option -> 'b option
 | 
|  |     25 |   val list_diff : ''a list -> ''a list -> ''a list
 | 
|  |     26 |   val list_prod : 'a list list -> 'a list -> 'a list -> 'a list list
 | 
|  |     27 |   val permute : ''a list -> ''a list list
 | 
|  |     28 |   val prefix_intersection_list :
 | 
|  |     29 |      ''a list -> ''a list -> ''a list
 | 
|  |     30 |   val repeat_until_fixpoint : (''a -> ''a) -> ''a -> ''a
 | 
|  |     31 |   val switch : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c
 | 
|  |     32 |   val zip_amap :
 | 
|  |     33 |        'a list ->
 | 
|  |     34 |        'b list ->
 | 
|  |     35 |        ('a * 'b) list -> ('a * 'b) list * ('a list * 'b list)
 | 
|  |     36 | 
 | 
|  |     37 |   val consts_in : term -> term list
 | 
| 59533 |     38 |   val head_quantified_variable : Proof.context -> int -> thm -> (string * typ) option
 | 
| 55596 |     39 |   val push_allvar_in : string -> term -> term
 | 
|  |     40 |   val strip_top_All_var : term -> (string * typ) * term
 | 
|  |     41 |   val strip_top_All_vars : term -> (string * typ) list * term
 | 
|  |     42 |   val strip_top_all_vars :
 | 
|  |     43 |      (string * typ) list -> term -> (string * typ) list * term
 | 
| 59533 |     44 |   val trace_tac' : Proof.context -> string ->
 | 
| 55596 |     45 |      ('a -> thm -> 'b Seq.seq) -> 'a -> thm -> 'b Seq.seq
 | 
|  |     46 |   val try_dest_Trueprop : term -> term
 | 
|  |     47 | 
 | 
|  |     48 |   val type_devar : ((indexname * sort) * typ) list -> term -> term
 | 
|  |     49 |   val diff_and_instantiate : Proof.context -> thm -> term -> term -> thm
 | 
|  |     50 | 
 | 
| 59533 |     51 |   val batter_tac : Proof.context -> int -> tactic
 | 
|  |     52 |   val break_hypotheses_tac : Proof.context -> int -> tactic
 | 
|  |     53 |   val clause_breaker_tac : Proof.context -> int -> tactic
 | 
| 55596 |     54 |   (* val dist_all_and_tac : Proof.context -> int -> tactic *)(*FIXME unused*)
 | 
|  |     55 |   val reassociate_conjs_tac : Proof.context -> int -> tactic
 | 
|  |     56 | 
 | 
|  |     57 |   val ASAP : (int -> tactic) -> (int -> tactic) -> int -> tactic
 | 
|  |     58 |   val COND' :
 | 
|  |     59 |      ('a -> thm -> bool) ->
 | 
|  |     60 |      ('a -> tactic) -> ('a -> tactic) -> 'a -> tactic
 | 
|  |     61 | 
 | 
|  |     62 |   val TERMFUN :
 | 
|  |     63 |      (term list * term -> 'a) -> int option -> thm -> 'a list
 | 
|  |     64 |   val TERMPRED :
 | 
|  |     65 |      (term -> bool) ->
 | 
|  |     66 |      (term -> bool) -> int option -> thm -> bool
 | 
|  |     67 | 
 | 
|  |     68 |   val guided_abstract :
 | 
|  |     69 |      bool -> term -> term -> ((string * typ) * term) * term list
 | 
|  |     70 |   val abstract :
 | 
|  |     71 |      term list -> term -> ((string * typ) * term) list * term
 | 
|  |     72 | end
 | 
|  |     73 | 
 | 
|  |     74 | structure TPTP_Reconstruct_Library : TPTP_RECONSTRUCT_LIBRARY =
 | 
|  |     75 | struct
 | 
|  |     76 | 
 | 
|  |     77 | (*zip as much as possible*)
 | 
|  |     78 | fun zip_amap [] ys acc = (acc, ([], ys))
 | 
|  |     79 |   | zip_amap xs [] acc = (acc, (xs, []))
 | 
|  |     80 |   | zip_amap (x :: xs) (y :: ys) acc =
 | 
|  |     81 |       zip_amap xs ys ((x, y) :: acc);
 | 
|  |     82 | 
 | 
|  |     83 | (*Pair a list up with the position number of each element,
 | 
|  |     84 |   starting from n*)
 | 
|  |     85 | fun enumerate n ls =
 | 
|  |     86 |   let
 | 
|  |     87 |     fun enumerate' [] _ acc = acc
 | 
|  |     88 |       | enumerate' (x :: xs) n acc = enumerate' xs (n + 1) ((n, x) :: acc)
 | 
|  |     89 |   in
 | 
|  |     90 |     enumerate' ls n []
 | 
|  |     91 |     |> rev
 | 
|  |     92 |   end
 | 
|  |     93 | 
 | 
|  |     94 | (*
 | 
|  |     95 | enumerate 0 [];
 | 
|  |     96 | enumerate 0 ["a", "b", "c"];
 | 
|  |     97 | *)
 | 
|  |     98 | 
 | 
|  |     99 | (*List subtraction*)
 | 
|  |    100 | fun list_diff l1 l2 =
 | 
| 58412 |    101 |   filter (fn x => forall (fn y => x <> y) l2) l1
 | 
| 55596 |    102 | 
 | 
|  |    103 | val _ = @{assert}
 | 
|  |    104 |   (list_diff [1,2,3] [2,4] = [1, 3])
 | 
|  |    105 | 
 | 
|  |    106 | (* [a,b] times_list [c,d] gives [[a,c,d], [b,c,d]] *)
 | 
|  |    107 | fun list_prod acc [] _ = rev acc
 | 
|  |    108 |   | list_prod acc (x :: xs) ys = list_prod ((x :: ys) :: acc) xs ys
 | 
|  |    109 | 
 | 
|  |    110 | fun repeat_until_fixpoint f x =
 | 
|  |    111 |   let
 | 
|  |    112 |     val x' = f x
 | 
|  |    113 |   in
 | 
|  |    114 |     if x = x' then x else repeat_until_fixpoint f x'
 | 
|  |    115 |   end
 | 
|  |    116 | 
 | 
|  |    117 | (*compute all permutations of a list*)
 | 
|  |    118 | fun permute l =
 | 
|  |    119 |   let
 | 
|  |    120 |     fun permute' (l, []) = [(l, [])]
 | 
|  |    121 |       | permute' (l, xs) =
 | 
|  |    122 |           map (fn x => (x :: l, filter (fn y => y <> x) xs)) xs
 | 
| 58411 |    123 |           |> maps permute'
 | 
| 55596 |    124 |   in
 | 
|  |    125 |     permute' ([], l)
 | 
|  |    126 |     |> map fst
 | 
|  |    127 |   end
 | 
|  |    128 | (*
 | 
|  |    129 | permute [1,2,3];
 | 
|  |    130 | permute ["A", "B"]
 | 
|  |    131 | *)
 | 
|  |    132 | 
 | 
|  |    133 | (*this exception is raised when the pair we wish to displace
 | 
|  |    134 |   isn't found in the association list*)
 | 
|  |    135 | exception DISPLACE_KV;
 | 
|  |    136 | (*move a key-value pair, determined by the k, to the beginning of
 | 
|  |    137 |   an association list. it moves the first occurrence of a pair
 | 
|  |    138 |   keyed by "k"*)
 | 
|  |    139 | local
 | 
|  |    140 |   fun fold_fun k (kv as (k', v)) (l, buff) =
 | 
|  |    141 |     if is_some buff then (kv :: l, buff)
 | 
|  |    142 |     else
 | 
|  |    143 |       if k = k' then
 | 
|  |    144 |         (l, SOME kv)
 | 
|  |    145 |       else
 | 
|  |    146 |         (kv :: l, buff)
 | 
|  |    147 | in
 | 
|  |    148 |   (*"k" is the key value of the pair we wish to displace*)
 | 
|  |    149 |   fun displace_kv k alist =
 | 
|  |    150 |     let
 | 
|  |    151 |       val (pre_alist, kv) = fold (fold_fun k) alist ([], NONE)
 | 
|  |    152 |     in
 | 
|  |    153 |       if is_some kv then
 | 
|  |    154 |         the kv :: rev pre_alist
 | 
|  |    155 |       else raise DISPLACE_KV
 | 
|  |    156 |     end
 | 
|  |    157 | end
 | 
|  |    158 | 
 | 
|  |    159 | (*Given two lists, it generates a new list where
 | 
|  |    160 |   the intersection of the lists forms the prefix
 | 
|  |    161 |   of the new list.*)
 | 
|  |    162 | local
 | 
|  |    163 |   fun prefix_intersection_list' (acc_pre, acc_pro) l1 l2 =
 | 
|  |    164 |     if null l1 then
 | 
|  |    165 |       List.rev acc_pre @ List.rev acc_pro
 | 
|  |    166 |     else if null l2 then
 | 
|  |    167 |       List.rev acc_pre @ l1 @ List.rev acc_pro
 | 
|  |    168 |     else
 | 
|  |    169 |       let val l1_hd = hd l1
 | 
|  |    170 |       in
 | 
|  |    171 |         prefix_intersection_list'
 | 
|  |    172 |          (if member (op =) l2 l1_hd then
 | 
|  |    173 |             (l1_hd :: acc_pre, acc_pro)
 | 
|  |    174 |           else
 | 
|  |    175 |            (acc_pre, l1_hd :: acc_pro))
 | 
|  |    176 |          (tl l1) l2
 | 
|  |    177 |       end
 | 
|  |    178 | in
 | 
|  |    179 |   fun prefix_intersection_list l1 l2 = prefix_intersection_list' ([], []) l1 l2
 | 
|  |    180 | end;
 | 
|  |    181 | 
 | 
|  |    182 | val _ = @{assert}
 | 
|  |    183 |   (prefix_intersection_list [1,2,3,4,5] [1,3,5] = [1, 3, 5, 2, 4]);
 | 
|  |    184 | 
 | 
|  |    185 | val _ = @{assert}
 | 
|  |    186 |   (prefix_intersection_list [1,2,3,4,5] [] = [1,2,3,4,5]);
 | 
|  |    187 | 
 | 
|  |    188 | val _ = @{assert}
 | 
|  |    189 |   (prefix_intersection_list [] [1,3,5] = [])
 | 
|  |    190 | 
 | 
|  |    191 | fun switch f y x = f x y
 | 
|  |    192 | 
 | 
|  |    193 | (*Given a value of type "'a option list", produce
 | 
|  |    194 |   a value of type "'a list" by dropping the NONE elements
 | 
|  |    195 |   and projecting the SOME elements.*)
 | 
|  |    196 | fun fold_options opt_list =
 | 
|  |    197 |   fold
 | 
|  |    198 |    (fn x => fn l => if is_some x then the x :: l else l)
 | 
|  |    199 |    opt_list
 | 
|  |    200 |    [];
 | 
|  |    201 | 
 | 
|  |    202 | val _ = @{assert}
 | 
|  |    203 |   ([2,0,1] =
 | 
|  |    204 |    fold_options [NONE, SOME 1, NONE, SOME 0, NONE, NONE, SOME 2]);
 | 
|  |    205 | 
 | 
|  |    206 | fun lift_option (f : 'a -> 'b) (x_opt : 'a option) : 'b option =
 | 
|  |    207 |   case x_opt of
 | 
|  |    208 |       NONE => NONE
 | 
|  |    209 |     | SOME x => SOME (f x)
 | 
|  |    210 | 
 | 
|  |    211 | fun break_seq x = (Seq.hd x, Seq.tl x)
 | 
|  |    212 | 
 | 
|  |    213 | exception BREAK_LIST
 | 
|  |    214 | fun break_list (x :: xs) = (x, xs)
 | 
|  |    215 |   | break_list _ = raise BREAK_LIST
 | 
|  |    216 | 
 | 
|  |    217 | exception MULTI_ELEMENT_LIST
 | 
|  |    218 | (*Try a number of predicates, in order, to find a single element.
 | 
|  |    219 |   Predicates are expected to either return an empty list or a
 | 
|  |    220 |   singleton list. If strict=true and list has more than one element,
 | 
|  |    221 |   then raise an exception. Otherwise try a new predicate.*)
 | 
|  |    222 | fun cascaded_filter_single strict preds l =
 | 
|  |    223 |   case preds of
 | 
|  |    224 |       [] => NONE
 | 
|  |    225 |     | (p :: ps) =>
 | 
|  |    226 |       case p l of
 | 
|  |    227 |           [] => cascaded_filter_single strict ps l
 | 
|  |    228 |         | [x] => SOME x
 | 
|  |    229 |         | l =>
 | 
|  |    230 |             if strict then raise MULTI_ELEMENT_LIST
 | 
|  |    231 |             else cascaded_filter_single strict ps l
 | 
|  |    232 | 
 | 
|  |    233 | (*concat but with optional before-and-after delimiters*)
 | 
|  |    234 | fun concat_between [] _ = []
 | 
|  |    235 |   | concat_between [l] _ = l
 | 
|  |    236 |   | concat_between (l :: ls) (seps as (bef, aft)) =
 | 
|  |    237 |     let
 | 
|  |    238 |       val pre = if is_some bef then the bef :: l else l
 | 
|  |    239 |       val mid = if is_some aft then [the aft] else []
 | 
|  |    240 |       val post = concat_between ls seps
 | 
|  |    241 |     in
 | 
|  |    242 |       pre @ mid @ post
 | 
|  |    243 |     end
 | 
|  |    244 | 
 | 
|  |    245 | (*Given a list, find an element satisfying pred, and return
 | 
|  |    246 |   a pair consisting of that element and the list minus the element.*)
 | 
|  |    247 | fun find_and_remove pred l =
 | 
|  |    248 |   find_index pred l
 | 
|  |    249 |   |> switch chop l
 | 
|  |    250 |   |> apsnd break_list
 | 
|  |    251 |   |> (fn (xs, (y, ys)) => (y, xs @ ys))
 | 
|  |    252 | 
 | 
|  |    253 | val _ = @{assert} (find_and_remove (curry (op =) 3) [0,1,2,3,4,5] = (3, [0,1,2,4,5]))
 | 
|  |    254 | 
 | 
|  |    255 | 
 | 
|  |    256 | (** Functions on terms **)
 | 
|  |    257 | 
 | 
|  |    258 | (*Extract the forall-prefix of a term, and return a pair consisting of the prefix
 | 
|  |    259 |   and the body*)
 | 
|  |    260 | local
 | 
|  |    261 |   (*Strip off HOL's All combinator if it's at the toplevel*)
 | 
|  |    262 |   fun try_dest_All (Const (@{const_name HOL.All}, _) $ t) = t
 | 
|  |    263 |     | try_dest_All (Const (@{const_name HOL.Trueprop}, _) $ t) = try_dest_All t
 | 
|  |    264 |     | try_dest_All t = t
 | 
|  |    265 | 
 | 
|  |    266 |   val _ = @{assert}
 | 
|  |    267 |     ((@{term "! x. (! y. P) = True"}
 | 
|  |    268 |       |> try_dest_All
 | 
|  |    269 |       |> Term.strip_abs_vars)
 | 
|  |    270 |      = [("x", @{typ "'a"})])
 | 
|  |    271 | 
 | 
|  |    272 |   val _ = @{assert}
 | 
|  |    273 |     ((@{prop "! x. (! y. P) = True"}
 | 
|  |    274 |       |> try_dest_All
 | 
|  |    275 |       |> Term.strip_abs_vars)
 | 
|  |    276 |      = [("x", @{typ "'a"})])
 | 
|  |    277 | 
 | 
|  |    278 |   fun strip_top_All_vars' once acc t =
 | 
|  |    279 |     let
 | 
|  |    280 |       val t' = try_dest_All t
 | 
|  |    281 |       val var =
 | 
|  |    282 |         try (Term.strip_abs_vars #> hd) t'
 | 
|  |    283 | 
 | 
|  |    284 |       fun strip v t =
 | 
|  |    285 |         (v, subst_bounds ([Free v], Term.strip_abs_body t))
 | 
|  |    286 |     in
 | 
|  |    287 |       if t' = t orelse is_none var then (acc, t)
 | 
|  |    288 |       else
 | 
|  |    289 |         let
 | 
|  |    290 |           val (v, t) = strip (the var) t'
 | 
|  |    291 |           val acc' = v :: acc
 | 
|  |    292 |         in
 | 
|  |    293 |           if once then (acc', t)
 | 
|  |    294 |           else strip_top_All_vars' once acc' t
 | 
|  |    295 |         end
 | 
|  |    296 |     end
 | 
|  |    297 | in
 | 
|  |    298 |   fun strip_top_All_vars t = strip_top_All_vars' false [] t
 | 
|  |    299 | 
 | 
|  |    300 | val _ =
 | 
|  |    301 |   let
 | 
|  |    302 |     val answer =
 | 
|  |    303 |       ([("x", @{typ "'a"})],
 | 
|  |    304 |        HOLogic.all_const @{typ "'a"} $
 | 
|  |    305 |         (HOLogic.eq_const @{typ "'a"} $
 | 
|  |    306 |          Free ("x", @{typ "'a"})))
 | 
|  |    307 |   in
 | 
|  |    308 |     @{assert}
 | 
|  |    309 |       ((@{term "! x. All (op = x)"}
 | 
|  |    310 |         |> strip_top_All_vars)
 | 
|  |    311 |        = answer)
 | 
|  |    312 |   end
 | 
|  |    313 | 
 | 
|  |    314 |   (*like strip_top_All_vars, but peels a single variable off, instead of all of them*)
 | 
|  |    315 |   fun strip_top_All_var t =
 | 
|  |    316 |     strip_top_All_vars' true [] t
 | 
|  |    317 |     |> apfst the_single
 | 
|  |    318 | end
 | 
|  |    319 | 
 | 
| 56245 |    320 | (*like strip_top_All_vars but for "Pure.all" instead of "HOL.All"*)
 | 
| 55596 |    321 | fun strip_top_all_vars acc t =
 | 
|  |    322 |   if Logic.is_all t then
 | 
|  |    323 |     let
 | 
|  |    324 |       val (v, t') = Logic.dest_all t
 | 
|  |    325 |       (*bound instances in t' are replaced with free vars*)
 | 
|  |    326 |     in
 | 
|  |    327 |       strip_top_all_vars (v :: acc) t'
 | 
|  |    328 |     end
 | 
|  |    329 |   else (acc, (*variables are returned in FILO order*)
 | 
|  |    330 |         t)
 | 
|  |    331 | 
 | 
|  |    332 | (*given a term "t"
 | 
|  |    333 |     ! X Y Z. t'
 | 
|  |    334 |   then then "push_allvar_in "X" t" will give
 | 
|  |    335 |     ! Y Z X. t'
 | 
|  |    336 | *)
 | 
|  |    337 | fun push_allvar_in v t =
 | 
|  |    338 |   let
 | 
|  |    339 |     val (vs, t') = strip_top_All_vars t
 | 
|  |    340 |     val vs' = displace_kv v vs
 | 
|  |    341 |   in
 | 
|  |    342 |     fold (fn (v, ty) => fn t =>
 | 
|  |    343 |       HOLogic.mk_all (v, ty, t)) vs' t'
 | 
|  |    344 |   end
 | 
|  |    345 | 
 | 
|  |    346 | (*Lists all consts in a term, uniquely*)
 | 
|  |    347 | fun consts_in (Const c) = [Const c]
 | 
|  |    348 |   | consts_in (Free _) = []
 | 
|  |    349 |   | consts_in (Var _) = []
 | 
|  |    350 |   | consts_in (Bound _) = []
 | 
|  |    351 |   | consts_in (Abs (_, _, t)) = consts_in t
 | 
|  |    352 |   | consts_in (t1 $ t2) = union (op =) (consts_in t1) (consts_in t2);
 | 
|  |    353 | 
 | 
|  |    354 | exception DIFF of term * term
 | 
|  |    355 | exception DIFF_TYPE of typ * typ
 | 
|  |    356 | (*This carries out naive form of matching.  It "diffs" two formulas,
 | 
|  |    357 |   to create a function which maps (schematic or non-schematic)
 | 
|  |    358 |   variables to terms.  The first argument is the more "general" term.
 | 
|  |    359 |   The second argument is used to find the "image" for the variables in
 | 
|  |    360 |   the first argument which don't appear in the second argument.
 | 
|  |    361 | 
 | 
|  |    362 |   Note that the list that is returned might have duplicate entries.
 | 
|  |    363 |   It's not checked to see if the same variable maps to different
 | 
|  |    364 |   values -- that should be regarded as an error.*)
 | 
|  |    365 | fun diff thy (initial as (t_gen, t)) =
 | 
|  |    366 |   let
 | 
|  |    367 |     fun diff_ty acc [] = acc
 | 
|  |    368 |       | diff_ty acc ((pair as (ty_gen, ty)) :: ts) =
 | 
|  |    369 |           case pair of
 | 
|  |    370 |               (Type (s1, ty_gens1), Type (s2, ty_gens2)) =>
 | 
|  |    371 |                 if s1 <> s2 orelse
 | 
|  |    372 |                  length ty_gens1 <> length ty_gens2 then
 | 
|  |    373 |                   raise (DIFF (t_gen, t))
 | 
|  |    374 |                 else
 | 
|  |    375 |                   diff_ty acc
 | 
|  |    376 |                    (ts @ ListPair.zip (ty_gens1, ty_gens2))
 | 
|  |    377 |             | (TFree (s1, so1), TFree (s2, so2)) =>
 | 
|  |    378 |                 if s1 <> s2 orelse
 | 
|  |    379 |                  not (Sign.subsort thy (so2, so1)) then
 | 
|  |    380 |                   raise (DIFF (t_gen, t))
 | 
|  |    381 |                 else
 | 
|  |    382 |                   diff_ty acc ts
 | 
|  |    383 |             | (TVar (idx1, so1), TVar (idx2, so2)) =>
 | 
|  |    384 |                 if idx1 <> idx2 orelse
 | 
|  |    385 |                  not (Sign.subsort thy (so2, so1)) then
 | 
|  |    386 |                   raise (DIFF (t_gen, t))
 | 
|  |    387 |                 else
 | 
|  |    388 |                   diff_ty acc ts
 | 
|  |    389 |             | (TFree _, _) => diff_ty (pair :: acc) ts
 | 
|  |    390 |             | (TVar _, _) => diff_ty (pair :: acc) ts
 | 
|  |    391 |             | _ => raise (DIFF_TYPE pair)
 | 
|  |    392 | 
 | 
|  |    393 |     fun diff' (acc as (acc_t, acc_ty)) (pair as (t_gen, t)) ts =
 | 
|  |    394 |       case pair of
 | 
|  |    395 |           (Const (s1, ty1), Const (s2, ty2)) =>
 | 
|  |    396 |             if s1 <> s2 orelse
 | 
|  |    397 |              not (Sign.typ_instance thy (ty2, ty1)) then
 | 
|  |    398 |               raise (DIFF (t_gen, t))
 | 
|  |    399 |             else
 | 
|  |    400 |               diff_probs acc ts
 | 
|  |    401 |         | (Free (s1, ty1), Free (s2, ty2)) =>
 | 
|  |    402 |             if s1 <> s2 orelse
 | 
|  |    403 |              not (Sign.typ_instance thy (ty2, ty1)) then
 | 
|  |    404 |               raise (DIFF (t_gen, t))
 | 
|  |    405 |             else
 | 
|  |    406 |               diff_probs acc ts
 | 
|  |    407 |         | (Var (idx1, ty1), Var (idx2, ty2)) =>
 | 
|  |    408 |             if idx1 <> idx2 orelse
 | 
|  |    409 |              not (Sign.typ_instance thy (ty2, ty1)) then
 | 
|  |    410 |               raise (DIFF (t_gen, t))
 | 
|  |    411 |             else
 | 
|  |    412 |               diff_probs acc ts
 | 
|  |    413 |         | (Bound i1, Bound i2) =>
 | 
|  |    414 |             if i1 <> i2 then
 | 
|  |    415 |               raise (DIFF (t_gen, t))
 | 
|  |    416 |             else
 | 
|  |    417 |               diff_probs acc ts
 | 
|  |    418 |         | (Abs (s1, ty1, t1), Abs (s2, ty2, t2)) =>
 | 
|  |    419 |             if s1 <> s2 orelse
 | 
|  |    420 |              not (Sign.typ_instance thy (ty2, ty1)) then
 | 
|  |    421 |               raise (DIFF (t_gen, t))
 | 
|  |    422 |             else
 | 
|  |    423 |               diff' acc (t1, t2) ts
 | 
|  |    424 |         | (ta1 $ ta2, tb1 $ tb2) =>
 | 
|  |    425 |             diff_probs acc ((ta1, tb1) :: (ta2, tb2) :: ts)
 | 
|  |    426 | 
 | 
|  |    427 |         (*the particularly important bit*)
 | 
|  |    428 |         | (Free (_, ty), _) =>
 | 
|  |    429 |             diff_probs
 | 
|  |    430 |              (pair :: acc_t,
 | 
|  |    431 |               diff_ty acc_ty [(ty, Term.fastype_of t)])
 | 
|  |    432 |              ts
 | 
|  |    433 |         | (Var (_, ty), _) =>
 | 
|  |    434 |             diff_probs
 | 
|  |    435 |              (pair :: acc_t,
 | 
|  |    436 |               diff_ty acc_ty [(ty, Term.fastype_of t)])
 | 
|  |    437 |              ts
 | 
|  |    438 | 
 | 
|  |    439 |         (*everything else is problematic*)
 | 
|  |    440 |         | _ => raise (DIFF (t_gen, t))
 | 
|  |    441 | 
 | 
|  |    442 |     and diff_probs acc ts =
 | 
|  |    443 |       case ts of
 | 
|  |    444 |           [] => acc
 | 
|  |    445 |         | (pair :: ts') => diff' acc pair ts'
 | 
|  |    446 |   in
 | 
|  |    447 |     diff_probs ([], []) [initial]
 | 
|  |    448 |   end
 | 
|  |    449 | 
 | 
|  |    450 | (*Abstracts occurrences of "t_sub" in "t", returning a list of
 | 
|  |    451 |   abstractions of "t" with a Var at each occurrence of "t_sub".
 | 
|  |    452 |   If "strong=true" then it uses strong abstraction (i.e., replaces
 | 
|  |    453 |    all occurrnces of "t_sub"), otherwise it uses weak abstraction
 | 
|  |    454 |    (i.e., replaces the occurrences one at a time).
 | 
|  |    455 |   NOTE there are many more possibilities between strong and week.
 | 
|  |    456 |     These can be enumerated by abstracting based on the powerset
 | 
|  |    457 |     of occurrences (minus the null element, which would correspond
 | 
|  |    458 |     to "t").
 | 
|  |    459 | *)
 | 
|  |    460 | fun guided_abstract strong t_sub t =
 | 
|  |    461 |   let
 | 
|  |    462 |     val varnames = Term.add_frees t [] |> map #1
 | 
|  |    463 |     val prefixK = "v"
 | 
|  |    464 |     val freshvar =
 | 
|  |    465 |       let
 | 
|  |    466 |         fun find_fresh i =
 | 
|  |    467 |           let
 | 
|  |    468 |             val varname = prefixK ^ Int.toString i
 | 
|  |    469 |           in
 | 
|  |    470 |             if member (op =) varnames varname then
 | 
|  |    471 |               find_fresh (i + 1)
 | 
|  |    472 |             else
 | 
|  |    473 |               (varname, fastype_of t_sub)
 | 
|  |    474 |           end
 | 
|  |    475 |       in
 | 
|  |    476 |         find_fresh 0
 | 
|  |    477 |       end
 | 
|  |    478 | 
 | 
|  |    479 |     fun guided_abstract' t =
 | 
|  |    480 |       case t of
 | 
|  |    481 |           Abs (s, ty, t') =>
 | 
|  |    482 |             if t = t_sub then [Free freshvar]
 | 
|  |    483 |             else
 | 
|  |    484 |               (map (fn t' => Abs (s, ty, t'))
 | 
|  |    485 |                (guided_abstract' t'))
 | 
|  |    486 |         | t1 $ t2 =>
 | 
|  |    487 |             if t = t_sub then [Free freshvar]
 | 
|  |    488 |             else
 | 
|  |    489 |                 (map (fn t' => t' $ t2)
 | 
|  |    490 |                   (guided_abstract' t1)) @
 | 
|  |    491 |                 (map (fn t' => t1 $ t')
 | 
|  |    492 |                   (guided_abstract' t2))
 | 
|  |    493 |         | _ =>
 | 
|  |    494 |             if t = t_sub then [Free freshvar]
 | 
|  |    495 |             else [t]
 | 
|  |    496 | 
 | 
|  |    497 |     fun guided_abstract_strong' t =
 | 
|  |    498 |       let
 | 
|  |    499 |         fun continue t = guided_abstract_strong' t
 | 
|  |    500 |           |> (fn x => if null x then t
 | 
|  |    501 |                 else the_single x)
 | 
|  |    502 |       in
 | 
|  |    503 |         case t of
 | 
|  |    504 |             Abs (s, ty, t') =>
 | 
|  |    505 |               if t = t_sub then [Free freshvar]
 | 
|  |    506 |               else
 | 
|  |    507 |                 [Abs (s, ty, continue t')]
 | 
|  |    508 |           | t1 $ t2 =>
 | 
|  |    509 |               if t = t_sub then [Free freshvar]
 | 
|  |    510 |               else
 | 
|  |    511 |                 [continue t1 $ continue t2]
 | 
|  |    512 |           | _ =>
 | 
|  |    513 |               if t = t_sub then [Free freshvar]
 | 
|  |    514 |               else [t]
 | 
|  |    515 |       end
 | 
|  |    516 | 
 | 
|  |    517 |   in
 | 
|  |    518 |     ((freshvar, t_sub),
 | 
|  |    519 |      if strong then guided_abstract_strong' t
 | 
|  |    520 |      else guided_abstract' t)
 | 
|  |    521 |   end
 | 
|  |    522 | 
 | 
|  |    523 | (*Carries out strong abstraction of a term guided by a list of
 | 
|  |    524 |   other terms.
 | 
|  |    525 |   In case some of the latter terms happen to be the same, it
 | 
|  |    526 |   only abstracts them once.
 | 
|  |    527 |   It returns the abstracted term, together with a map from
 | 
|  |    528 |    the fresh names to the terms.*)
 | 
|  |    529 | fun abstract ts t =
 | 
|  |    530 |   fold_map (apsnd the_single oo (guided_abstract true)) ts t
 | 
|  |    531 |   |> (fn (v_and_ts, t') =>
 | 
|  |    532 |        let
 | 
|  |    533 |          val (vs, ts) = ListPair.unzip v_and_ts
 | 
|  |    534 |          val vs' =
 | 
|  |    535 |            (* list_diff vs (list_diff (Term.add_frees t' []) vs) *)
 | 
|  |    536 |            Term.add_frees t' []
 | 
|  |    537 |            |> list_diff vs
 | 
|  |    538 |            |> list_diff vs
 | 
|  |    539 |          val v'_and_ts =
 | 
|  |    540 |            map (fn v =>
 | 
|  |    541 |              (v, AList.lookup (op =) v_and_ts v |> the))
 | 
|  |    542 |             vs'
 | 
|  |    543 |        in
 | 
|  |    544 |          (v'_and_ts, t')
 | 
|  |    545 |        end)
 | 
|  |    546 | 
 | 
|  |    547 | (*Instantiate type variables in a term, based on a type environment*)
 | 
|  |    548 | fun type_devar (tyenv : ((indexname * sort) * typ) list) (t : term) : term =
 | 
|  |    549 |   case t of
 | 
|  |    550 |       Const (s, ty) => Const (s, Term_Subst.instantiateT tyenv ty)
 | 
|  |    551 |     | Free (s, ty) => Free (s, Term_Subst.instantiateT tyenv ty)
 | 
|  |    552 |     | Var (idx, ty) => Var (idx, Term_Subst.instantiateT tyenv ty)
 | 
|  |    553 |     | Bound _ => t
 | 
|  |    554 |     | Abs (s, ty, t') =>
 | 
|  |    555 |         Abs (s, Term_Subst.instantiateT tyenv ty, type_devar tyenv t')
 | 
|  |    556 |     | t1 $ t2 => type_devar tyenv t1 $ type_devar tyenv t2
 | 
|  |    557 | 
 | 
|  |    558 | (*Take a "diff" between an (abstract) thm's term, and another term
 | 
|  |    559 |   (the latter is an instance of the form), then instantiate the
 | 
|  |    560 |   abstract theorem. This is a way of turning the latter term into
 | 
|  |    561 |   a theorem, but without exposing the proof-search functions to
 | 
|  |    562 |   complex terms.
 | 
|  |    563 |   In addition to the abstract thm ("scheme_thm"), this function is
 | 
|  |    564 |   also supplied with the (sub)term of the abstract thm ("scheme_t")
 | 
|  |    565 |   we want to use in the diff, in case only part of "scheme_t"
 | 
| 59582 |    566 |   might be needed (not the whole "Thm.prop_of scheme_thm")*)
 | 
| 55596 |    567 | fun diff_and_instantiate ctxt scheme_thm scheme_t instance_t =
 | 
|  |    568 |   let
 | 
|  |    569 |     val (term_pairing, type_pairing) =
 | 
| 59639 |    570 |       diff (Proof_Context.theory_of ctxt) (scheme_t, instance_t)
 | 
| 55596 |    571 | 
 | 
|  |    572 |     (*valuation of type variables*)
 | 
| 59639 |    573 |     val typeval = map (apply2 (Thm.ctyp_of ctxt)) type_pairing
 | 
| 55596 |    574 | 
 | 
|  |    575 |     val typeval_env =
 | 
|  |    576 |       map (apfst dest_TVar) type_pairing
 | 
|  |    577 |     (*valuation of term variables*)
 | 
|  |    578 |     val termval =
 | 
|  |    579 |       map (apfst (type_devar typeval_env)) term_pairing
 | 
| 59639 |    580 |       |> map (apply2 (Thm.cterm_of ctxt))
 | 
| 55596 |    581 |   in
 | 
|  |    582 |     Thm.instantiate (typeval, termval) scheme_thm
 | 
|  |    583 |   end
 | 
|  |    584 | 
 | 
|  |    585 | (*FIXME this is bad form?*)
 | 
|  |    586 | val try_dest_Trueprop = perhaps (try HOLogic.dest_Trueprop)
 | 
|  |    587 | 
 | 
|  |    588 | 
 | 
|  |    589 | (** Some tacticals **)
 | 
|  |    590 | 
 | 
|  |    591 | (*Lift COND to be parametrised by subgoal number*)
 | 
|  |    592 | fun COND' sat' tac'1 tac'2 i =
 | 
|  |    593 |   COND (sat' i) (tac'1 i) (tac'2 i)
 | 
|  |    594 | 
 | 
|  |    595 | (*Apply simplification ("wittler") as few times as possible
 | 
|  |    596 |   before being able to apply a tactic ("tac").
 | 
|  |    597 |   This is like a lazy version of REPEAT, since it attempts
 | 
|  |    598 |   to REPEAT a tactic the smallest number times as possible,
 | 
|  |    599 |   to make some other tactic succeed subsequently.*)
 | 
|  |    600 | fun ASAP wittler (tac : int -> tactic) (i : int) = fn st =>
 | 
|  |    601 |   let
 | 
|  |    602 |     val tac_result = tac i st
 | 
|  |    603 |     val pulled_tac_result = Seq.pull tac_result
 | 
|  |    604 |     val tac_failed =
 | 
|  |    605 |       is_none pulled_tac_result orelse
 | 
|  |    606 |        not (has_fewer_prems 1 (fst (the pulled_tac_result)))
 | 
|  |    607 |   in
 | 
|  |    608 |     if tac_failed then (wittler THEN' ASAP wittler tac) i st
 | 
|  |    609 |     else tac_result
 | 
|  |    610 |   end
 | 
|  |    611 | 
 | 
|  |    612 | 
 | 
|  |    613 | (** Some tactics **)
 | 
|  |    614 | 
 | 
| 59533 |    615 | fun break_hypotheses_tac ctxt =
 | 
|  |    616 |   CHANGED o
 | 
|  |    617 |    ((REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}) THEN'
 | 
|  |    618 |     (REPEAT_DETERM o eresolve_tac ctxt @{thms disjE}))
 | 
| 55596 |    619 | 
 | 
|  |    620 | (*Prove subgoals of form A ==> B1 | ... | A | ... | Bn*)
 | 
| 59533 |    621 | fun clause_breaker_tac ctxt =
 | 
|  |    622 |   (REPEAT o resolve_tac ctxt @{thms disjI1 disjI2 conjI}) THEN'
 | 
|  |    623 |   assume_tac ctxt
 | 
| 55596 |    624 | 
 | 
|  |    625 | (*
 | 
|  |    626 |   Refines a subgoal have the form:
 | 
|  |    627 |     A1 ... An ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
 | 
|  |    628 |   into multiple subgoals of the form:
 | 
|  |    629 |     A'1 ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
 | 
|  |    630 |      :
 | 
|  |    631 |     A'm ==> B1 | ... | Aj | ... | Bi | ... | Ak | ...
 | 
|  |    632 |   where {A'1 .. A'm} is disjoint from {B1, ..., Aj, ..., Bi, ..., Ak, ...}
 | 
|  |    633 |   (and solves the subgoal completely if the first set is empty)
 | 
|  |    634 | *)
 | 
| 59533 |    635 | fun batter_tac ctxt i =
 | 
|  |    636 |   break_hypotheses_tac ctxt i THEN
 | 
|  |    637 |   ALLGOALS (TRY o clause_breaker_tac ctxt)
 | 
| 55596 |    638 | 
 | 
|  |    639 | (*Same idiom as ex_expander_tac*)
 | 
|  |    640 | fun dist_all_and_tac ctxt i =
 | 
|  |    641 |    let
 | 
|  |    642 |      val simpset =
 | 
|  |    643 |        empty_simpset ctxt
 | 
|  |    644 |        |> Simplifier.add_simp
 | 
|  |    645 |            @{lemma "! x. P x & Q x \<equiv> (! x. P x) & (! x. Q x)"
 | 
|  |    646 |               by (rule eq_reflection, auto)}
 | 
|  |    647 |    in
 | 
|  |    648 |      CHANGED (asm_full_simp_tac simpset i)
 | 
|  |    649 |    end
 | 
|  |    650 | 
 | 
|  |    651 | fun reassociate_conjs_tac ctxt =
 | 
|  |    652 |   asm_full_simp_tac
 | 
|  |    653 |    (Simplifier.add_simp
 | 
|  |    654 |     @{lemma "(A & B) & C == A & B & C" by auto} (*FIXME duplicates @{thm simp_meta(3)}*)
 | 
|  |    655 |     (Raw_Simplifier.empty_simpset ctxt))
 | 
|  |    656 |   #> CHANGED
 | 
|  |    657 |   #> REPEAT_DETERM
 | 
|  |    658 | 
 | 
|  |    659 | 
 | 
|  |    660 | (** Subgoal analysis **)
 | 
|  |    661 | 
 | 
|  |    662 | (*Given an inference
 | 
|  |    663 |         C
 | 
|  |    664 |       -----
 | 
|  |    665 |         D
 | 
|  |    666 |   This function returns "SOME X" if C = "! X. C'".
 | 
|  |    667 |   If C has no quantification prefix, then returns NONE.*)
 | 
| 59533 |    668 | fun head_quantified_variable ctxt i = fn st =>
 | 
| 55596 |    669 |   let
 | 
|  |    670 |     val gls =
 | 
| 59582 |    671 |       Thm.prop_of st
 | 
| 55596 |    672 |       |> Logic.strip_horn
 | 
|  |    673 |       |> fst
 | 
|  |    674 | 
 | 
|  |    675 |     val hypos =
 | 
|  |    676 |       if null gls then []
 | 
|  |    677 |       else
 | 
|  |    678 |         rpair (i - 1) gls
 | 
|  |    679 |         |> uncurry nth
 | 
|  |    680 |         |> strip_top_all_vars []
 | 
|  |    681 |         |> snd
 | 
|  |    682 |         |> Logic.strip_horn
 | 
|  |    683 |         |> fst
 | 
|  |    684 | 
 | 
|  |    685 |     fun foralls_of_hd_hypos () =
 | 
|  |    686 |       hd hypos
 | 
|  |    687 |       |> try_dest_Trueprop
 | 
|  |    688 |       |> strip_top_All_vars
 | 
|  |    689 |       |> #1
 | 
|  |    690 |       |> rev
 | 
|  |    691 | 
 | 
|  |    692 |     val quantified_variables = foralls_of_hd_hypos ()
 | 
|  |    693 |   in
 | 
|  |    694 |     if null hypos orelse null quantified_variables then NONE
 | 
|  |    695 |     else SOME (hd quantified_variables)
 | 
|  |    696 |   end
 | 
|  |    697 | 
 | 
|  |    698 | 
 | 
|  |    699 | (** Builders for goal analysers or transformers **)
 | 
|  |    700 | 
 | 
|  |    701 | (*Lifts function over terms to apply it to subgoals.
 | 
|  |    702 |   "fun_over_terms" has type (term list * term -> 'a), where
 | 
|  |    703 |   (term list * term) will be the term representations of the
 | 
|  |    704 |   hypotheses and conclusion.
 | 
|  |    705 |   if i_opt=SOME i then applies fun_over_terms to that
 | 
|  |    706 |    subgoal and returns singleton result.
 | 
|  |    707 |   otherwise applies fun_over_terms to all subgoals and return
 | 
|  |    708 |    list of results.*)
 | 
|  |    709 | fun TERMFUN
 | 
|  |    710 |  (fun_over_terms : term list * term -> 'a)
 | 
|  |    711 |  (i_opt : int option) : thm -> 'a list = fn st =>
 | 
|  |    712 |   let
 | 
|  |    713 |     val t_raws =
 | 
| 59586 |    714 |         Thm.prop_of st
 | 
| 55596 |    715 |         |> strip_top_all_vars []
 | 
|  |    716 |         |> snd
 | 
|  |    717 |         |> Logic.strip_horn
 | 
|  |    718 |         |> fst
 | 
|  |    719 |   in
 | 
|  |    720 |     if null t_raws then []
 | 
|  |    721 |     else
 | 
|  |    722 |       let
 | 
|  |    723 |         val ts =
 | 
|  |    724 |           let
 | 
|  |    725 |             val stripper =
 | 
|  |    726 |               strip_top_all_vars []
 | 
|  |    727 |               #> snd
 | 
|  |    728 |               #> Logic.strip_horn
 | 
|  |    729 |               #> apsnd try_dest_Trueprop
 | 
|  |    730 |               #> apfst (map try_dest_Trueprop)
 | 
|  |    731 |           in
 | 
|  |    732 |             map stripper t_raws
 | 
|  |    733 |           end
 | 
|  |    734 |       in
 | 
|  |    735 |         case i_opt of
 | 
|  |    736 |             NONE =>
 | 
|  |    737 |               map fun_over_terms ts
 | 
|  |    738 |           | SOME i =>
 | 
|  |    739 |               nth ts (i - 1)
 | 
|  |    740 |               |> fun_over_terms
 | 
|  |    741 |               |> single
 | 
|  |    742 |       end
 | 
|  |    743 |   end
 | 
|  |    744 | 
 | 
|  |    745 | (*Applies a predicate to subgoal(s) conclusion(s)*)
 | 
|  |    746 | fun TERMPRED
 | 
|  |    747 |  (hyp_pred_over_terms : term -> bool)
 | 
|  |    748 |  (conc_pred_over_terms : term -> bool)
 | 
|  |    749 |  (i_opt : int option) : thm -> bool = fn st =>
 | 
|  |    750 |     let
 | 
|  |    751 |       val hyp_results =
 | 
|  |    752 |         TERMFUN (fst (*discard hypotheses*)
 | 
|  |    753 |                  #> map hyp_pred_over_terms) i_opt st
 | 
|  |    754 |       val conc_results =
 | 
|  |    755 |         TERMFUN (snd (*discard hypotheses*)
 | 
|  |    756 |                  #> conc_pred_over_terms) i_opt st
 | 
|  |    757 |       val _ = @{assert} (length hyp_results = length conc_results)
 | 
|  |    758 |     in
 | 
|  |    759 |       if null hyp_results then true
 | 
|  |    760 |       else
 | 
|  |    761 |         let
 | 
|  |    762 |           val hyps_conjoined =
 | 
|  |    763 |             fold (fn a => fn b =>
 | 
| 58412 |    764 |               b andalso (forall (fn x => x) a)) hyp_results true
 | 
| 55596 |    765 |           val concs_conjoined =
 | 
|  |    766 |             fold (fn a => fn b =>
 | 
|  |    767 |               b andalso a) conc_results true
 | 
|  |    768 |         in hyps_conjoined andalso concs_conjoined end
 | 
|  |    769 |     end
 | 
|  |    770 | 
 | 
|  |    771 | 
 | 
|  |    772 | (** Tracing **)
 | 
|  |    773 | (*If "tac i st" succeeds then msg is printed to "trace" channel*)
 | 
| 59533 |    774 | fun trace_tac' ctxt msg tac i st =
 | 
| 55596 |    775 |   let
 | 
|  |    776 |     val result = tac i st
 | 
|  |    777 |   in
 | 
|  |    778 |     if Config.get ctxt tptp_trace_reconstruction andalso
 | 
|  |    779 |      not (is_none (Seq.pull result)) then
 | 
|  |    780 |       (tracing msg; result)
 | 
|  |    781 |     else result
 | 
|  |    782 |   end
 | 
|  |    783 | 
 | 
|  |    784 | end
 |