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