src/Pure/IsaPlanner/rw_inst.ML
author wenzelm
Sat, 17 Sep 2005 18:11:28 +0200
changeset 17468 7c040a5fd171
parent 16179 fa7e70be26b0
child 18127 9f03d8a9a81b
permissions -rw-r--r--
pretty_thm_aux: aconv hyps;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
16179
fa7e70be26b0 header;
wenzelm
parents: 15959
diff changeset
     2
(*  Title:      Pure/IsaPlanner/rw_inst.ML
fa7e70be26b0 header;
wenzelm
parents: 15959
diff changeset
     3
    ID:         $Id$
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
                lucas.dixon@ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
    Created:    25 Aug 2004
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     7
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     8
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
    rewriting using a conditional meta-equality theorem which supports 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
    schematic variable instantiation.
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    13
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    14
*)   
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    15
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
signature RW_INST =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
sig
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    19
  (* Rewrite: give it instantiation infromation, a rule, and the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
  target thm, and it will return the rewritten target thm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
  val rw :
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    22
      ((Term.indexname * (Term.sort * Term.typ)) list *  (* type var instantiations *)
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    23
       (Term.indexname * (Term.typ * Term.term)) list)  (* schematic var instantiations *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    24
      * (string * Term.typ) list           (* Fake named bounds + types *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    25
      * (string * Term.typ) list           (* names of bound + types *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    26
      * Term.term ->                       (* outer term for instantiation *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    27
      Thm.thm ->                           (* rule with indexies lifted *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    28
      Thm.thm ->                           (* target thm *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    29
      Thm.thm                              (* rewritten theorem possibly 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    30
                                              with additional premises for 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    31
                                              rule conditions *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    32
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    33
  (* used tools *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    34
  val mk_abstractedrule :
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    35
      (string * Term.typ) list (* faked outer bound *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    36
      -> (string * Term.typ) list (* hopeful name of outer bounds *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    37
      -> Thm.thm -> Thm.cterm list * Thm.thm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    38
  val mk_fixtvar_tyinsts :
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    39
      (Term.indexname * (Term.sort * Term.typ)) list ->
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    40
      Term.term list -> ((string * int) * (Term.sort * Term.typ)) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    41
                        * (string * Term.sort) list
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    42
  val mk_renamings :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
      Term.term -> Thm.thm -> (((string * int) * Term.typ) * Term.term) list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
  val new_tfree :
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
      ((string * int) * Term.sort) *
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    46
      (((string * int) * (Term.sort * Term.typ)) list * string list) ->
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    47
      ((string * int) * (Term.sort * Term.typ)) list * string list
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    48
  val cross_inst : (Term.indexname * (Term.typ * Term.term)) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    49
                   -> (Term.indexname *(Term.typ * Term.term)) list
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    50
  val cross_inst_typs : (Term.indexname * (Term.sort * Term.typ)) list 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    51
                   -> (Term.indexname * (Term.sort * Term.typ)) list
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    53
  val beta_contract : Thm.thm -> Thm.thm
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    54
  val beta_eta_contract : Thm.thm -> Thm.thm
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    55
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    58
structure RWInst 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    59
(* : RW_INST *)
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    60
= struct
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
(* beta contract the theorem *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    64
fun beta_contract thm = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
    equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    66
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    67
(* beta-eta contract the theorem *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
    68
fun beta_eta_contract thm = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
      val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
      val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
    in thm3 end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    74
(* Given a list of variables that were bound, and a that has been
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    75
instantiated with free variable placeholders for the bound vars, it
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    76
creates an abstracted version of the theorem, with local bound vars as
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
lambda-params:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
Ts: 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
("x", ty)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    81
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    82
rule::
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    83
C :x ==> P :x = Q :x
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    84
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    85
results in:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    86
("!! x. C x", (%x. p x = %y. p y) [!! x. C x])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    87
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    88
note: assumes rule is instantiated
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    89
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    90
(* Note, we take abstraction in the order of last abstraction first *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    91
fun mk_abstractedrule TsFake Ts rule = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    92
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
      val ctermify = Thm.cterm_of (Thm.sign_of_thm rule);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    94
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    95
      (* now we change the names of temporary free vars that represent 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    96
         bound vars with binders outside the redex *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
      val prop = Thm.prop_of rule;
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
    98
      val names = TermLib.usednames_of_thm rule;
15560
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
    99
      val (fromnames,tonames,names2,Ts') = 
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   100
          Library.foldl (fn ((rnf,rnt,names, Ts''),((faken,_),(n,ty))) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
                    let val n2 = Term.variant names n in
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   102
                      (ctermify (Free(faken,ty)) :: rnf,
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   103
                       ctermify (Free(n2,ty)) :: rnt, 
15560
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
   104
                       n2 :: names,
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   105
                       (n2,ty) :: Ts'')
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   106
                    end)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   107
                (([],[],names, []), TsFake~~Ts);
15560
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
   108
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   109
      (* rename conflicting free's in the rule to avoid cconflicts
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   110
      with introduced vars from bounds outside in redex *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   111
      val rule' = rule |> Drule.forall_intr_list fromnames
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
                       |> Drule.forall_elim_list tonames;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   113
      
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   114
      (* make unconditional rule and prems *)
15560
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
   115
      val (uncond_rule, cprems) = IsaND.allify_conditions ctermify (rev Ts') 
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
   116
                                                          rule';
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   117
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
      (* using these names create lambda-abstracted version of the rule *)
15560
c862d556fb18 lucas - fixed bug with name capture variables bound outside redex could (previously)conflict with scheme variables that occur in the conditions of an equation, and which were renamed to avoid conflict with another instantiation. This has now been fixed.
dixon
parents: 15481
diff changeset
   119
      val abstractions = rev (Ts' ~~ tonames);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15560
diff changeset
   120
      val abstract_rule = Library.foldl (fn (th,((n,ty),ct)) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
                                    Thm.abstract_rule n ct th)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   122
                                (uncond_rule, abstractions);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   123
    in (cprems, abstract_rule) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   125
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   126
(* given names to avoid, and vars that need to be fixed, it gives
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   127
unique new names to the vars so that they can be fixed as free
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   128
variables *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   129
(* make fixed unique free variable instantiations for non-ground vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   130
(* Create a table of vars to be renamed after instantiation - ie
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   131
      other uninstantiated vars in the hyps of the rule 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   132
      ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   133
fun mk_renamings tgt rule_inst = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
    let
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   135
      val rule_conds = Thm.prems_of rule_inst
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   136
      val names = foldr Term.add_term_names [] (tgt :: rule_conds);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   137
      val (conds_tyvs,cond_vs) = 
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15560
diff changeset
   138
          Library.foldl (fn ((tyvs, vs), t) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   139
                    (Library.union
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   140
                       (Term.term_tvars t, tyvs),
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   141
                     Library.union 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   142
                       (map Term.dest_Var (Term.term_vars t), vs))) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   143
                (([],[]), rule_conds);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   144
      val termvars = map Term.dest_Var (Term.term_vars tgt); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   145
      val vars_to_fix = Library.union (termvars, cond_vs);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   146
      val (renamings, names2) = 
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   147
          foldr (fn (((n,i),ty), (vs, names')) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
                    let val n' = Term.variant names' n in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
                      ((((n,i),ty), Free (n', ty)) :: vs, n'::names')
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   150
                    end)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   151
                ([], names) vars_to_fix;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   152
    in renamings end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   153
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   154
(* make a new fresh typefree instantiation for the given tvar *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   155
fun new_tfree (tv as (ix,sort), (pairs,used)) =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   156
      let val v = variant used (string_of_indexname ix)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   157
      in  ((ix,(sort,TFree(v,sort)))::pairs, v::used)  end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   158
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   159
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   160
(* make instantiations to fix type variables that are not 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   161
   already instantiated (in ignore_ixs) from the list of terms. *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   162
fun mk_fixtvar_tyinsts ignore_insts ts = 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   163
    let 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   164
      val ignore_ixs = map fst ignore_insts;
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   165
      val (tvars, tfrees) = 
15798
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   166
            foldr (fn (t, (varixs, tfrees)) => 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   167
                      (Term.add_term_tvars (t,varixs),
15798
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   168
                       Term.add_term_tfrees (t,tfrees)))
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   169
                  ([],[]) ts;
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   170
        val unfixed_tvars = 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   171
            List.filter (fn (ix,s) => not (ix mem ignore_ixs)) tvars;
15798
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   172
        val (fixtyinsts, _) = foldr new_tfree ([], map fst tfrees) unfixed_tvars
016f3be5a5ec Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15574
diff changeset
   173
    in (fixtyinsts, tfrees) end;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   175
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   176
(* cross-instantiate the instantiations - ie for each instantiation
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   177
replace all occurances in other instantiations - no loops are possible
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   178
and thus only one-parsing of the instantiations is necessary. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
fun cross_inst insts = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   180
    let 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   181
      fun instL (ix, (ty,t)) = 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   182
          map (fn (ix2,(ty2,t2)) => 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   183
                  (ix2, (ty2,Term.subst_vars ([], [(ix, t)]) t2)));
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   185
      fun cross_instL ([], l) = rev l
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   186
        | cross_instL ((ix, t) :: insts, l) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
          cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
    in cross_instL (insts, []) end;
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   190
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   191
(* as above but for types -- I don't know if this is needed, will we ever incur mixed up types? *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   192
fun cross_inst_typs insts = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
    let 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   194
      fun instL (ix, (srt,ty)) = 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   195
          map (fn (ix2,(srt2,ty2)) => 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   196
                  (ix2, (srt2,Term.typ_subst_TVars [(ix, ty)] ty2)));
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   197
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   198
      fun cross_instL ([], l) = rev l
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   199
        | cross_instL ((ix, t) :: insts, l) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   200
          cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l));
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   201
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   202
    in cross_instL (insts, []) end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   203
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   204
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   205
(* assume that rule and target_thm have distinct var names. THINK:
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   206
efficient version with tables for vars for: target vars, introduced
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   207
vars, and rule vars, for quicker instantiation?  The outerterm defines
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   208
which part of the target_thm was modified.  Note: we take Ts in the
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   209
upterm order, ie last abstraction first., and with an outeterm where
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   210
the abstracted subterm has the arguments in the revered order, ie
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   211
first abstraction first.  FakeTs has abstractions using the fake name
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   212
- ie the name distinct from all other abstractions. *)
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   213
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15798
diff changeset
   214
fun rw ((nonfixed_typinsts, unprepinsts), FakeTs, Ts, outerterm) rule target_thm = 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   215
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   216
      (* general signature info *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   217
      val target_sign = (Thm.sign_of_thm target_thm);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   218
      val ctermify = Thm.cterm_of target_sign;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   219
      val ctypeify = Thm.ctyp_of target_sign;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   220
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   221
      (* fix all non-instantiated tvars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   222
      val (fixtyinsts, othertfrees) = 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   223
          mk_fixtvar_tyinsts nonfixed_typinsts
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   224
                             [Thm.prop_of rule, Thm.prop_of target_thm];
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   225
      val new_fixed_typs = map (fn ((s,i),(srt,ty)) => (Term.dest_TFree ty))
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   226
                               fixtyinsts;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   227
      val typinsts = cross_inst_typs (nonfixed_typinsts @ fixtyinsts);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   228
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   229
      (* certified instantiations for types *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   230
      val ctyp_insts = 
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   231
          map (fn (ix,(s,ty)) => (ctypeify (TVar (ix,s)), ctypeify ty)) 
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   232
              typinsts;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   233
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   234
      (* type instantiated versions *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   235
      val tgt_th_tyinst = Thm.instantiate (ctyp_insts,[]) target_thm;
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   236
      val rule_tyinst =  Thm.instantiate (ctyp_insts,[]) rule;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   237
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   238
      val term_typ_inst = map (fn (ix,(srt,ty)) => (ix,ty)) typinsts;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   239
      (* type instanitated outer term *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   240
      val outerterm_tyinst = Term.subst_TVars term_typ_inst outerterm;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   241
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   242
      val FakeTs_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) 
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   243
                              FakeTs;
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   244
      val Ts_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) 
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   245
                          Ts;
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   246
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   247
      (* type-instantiate the var instantiations *)
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   248
      val insts_tyinst = foldr (fn ((ix,(ty,t)),insts_tyinst) => 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   249
                            (ix, (Term.typ_subst_TVars term_typ_inst ty, 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   250
                                  Term.subst_TVars term_typ_inst t))
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   251
                            :: insts_tyinst)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   252
                        [] unprepinsts;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   253
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   254
      (* cross-instantiate *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   255
      val insts_tyinst_inst = cross_inst insts_tyinst;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   256
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   257
      (* create certms of instantiations *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   258
      val cinsts_tyinst = 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   259
          map (fn (ix,(ty,t)) => (ctermify (Var (ix, ty)), 
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   260
                                  ctermify t)) insts_tyinst_inst;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   261
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   262
      (* The instantiated rule *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   263
      val rule_inst = rule_tyinst |> Thm.instantiate ([], cinsts_tyinst);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   264
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   265
      (* Create a table of vars to be renamed after instantiation - ie
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   266
      other uninstantiated vars in the hyps the *instantiated* rule 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   267
      ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   268
      val renamings = mk_renamings (Thm.prop_of tgt_th_tyinst) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   269
                                   rule_inst;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   270
      val cterm_renamings = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   271
          map (fn (x,y) => (ctermify (Var x), ctermify y)) renamings;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   272
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   273
      (* Create the specific version of the rule for this target application *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   274
      val outerterm_inst = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   275
          outerterm_tyinst 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   276
            |> Term.subst_Vars (map (fn (ix,(ty,t)) => (ix,t)) insts_tyinst_inst)
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   277
            |> Term.subst_Vars (map (fn ((ix,ty),t) => (ix,t)) renamings);
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   278
      val couter_inst = Thm.reflexive (ctermify outerterm_inst);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   279
      val (cprems, abstract_rule_inst) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   280
          rule_inst |> Thm.instantiate ([], cterm_renamings)
15959
366d39e95d3c lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents: 15915
diff changeset
   281
                    |> mk_abstractedrule FakeTs_tyinst Ts_tyinst;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   282
      val specific_tgt_rule = 
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   283
          beta_eta_contract
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   284
            (Thm.combination couter_inst abstract_rule_inst);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   285
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   286
      (* create an instantiated version of the target thm *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   287
      val tgt_th_inst = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   288
          tgt_th_tyinst |> Thm.instantiate ([], cinsts_tyinst)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   289
                        |> Thm.instantiate ([], cterm_renamings);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   290
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   291
      val (vars,frees_of_fixed_vars) = Library.split_list cterm_renamings;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   292
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   293
    in
15915
b0e8b37642a4 lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents: 15814
diff changeset
   294
      (beta_eta_contract tgt_th_inst)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   295
        |> Thm.equal_elim specific_tgt_rule
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   296
        |> Drule.implies_intr_list cprems
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   297
        |> Drule.forall_intr_list frees_of_fixed_vars
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   298
        |> Drule.forall_elim_list vars
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   299
        |> fst o Thm.varifyT' othertfrees
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   300
        |> Drule.zero_var_indexes
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   301
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   302
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   303
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   304
end; (* struct *)