src/Pure/IsaPlanner/rw_tools.ML
author wenzelm
Tue, 17 May 2005 18:10:44 +0200
changeset 15991 670f8e4b5a98
parent 15814 d65f461c8672
child 16179 fa7e70be26b0
permissions -rw-r--r--
added read_variable: optional question mark on input; removed obsolete token_assoc; scan_indexname: improved treatment of \<^isub> and \<^isup>; read_var: more robust against bad input; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     2
(*  Title:      sys/rw_tools.ML
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
                lucas.dixon@ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
    Created:    28 May 2004
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
*)
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
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
    term related tools used for rewriting
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
*)   
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
signature RWTOOLS =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
sig
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
end;
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
structure RWTools 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
= struct
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    22
(* THINKABOUT: don't need this: should be able to generate the paired 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    23
   NsTs directly ? --already implemented as ~~
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    24
fun join_lists ([],[]) = []
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    25
  | join_lists (x::xs, y::ys) = (x,y) :: (join_lists (xs,ys))
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
  | join_lists (_, _) = raise ERROR_MESSAGE "join_lists: unequal size lists";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
 *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    28
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    29
(* fake free variable names for locally bound variables - these work
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    30
as placeholders. *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    31
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    32
(* don't use dest_fake.. - we should instead be working with numbers
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    33
and a list... else we rely on naming conventions which can break, or
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    34
be violated - in contrast list locations are correct by
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    35
construction/definition. *)
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    36
(*
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    37
fun dest_fake_bound_name n = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    38
    case (explode n) of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    39
      (":" :: realchars) => implode realchars
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    40
    | _ => n; *)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
fun is_fake_bound_name n = (hd (explode n) = ":");
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    42
fun mk_fake_bound_name n = ":b_" ^ n;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    46
(* fake free variable names for local meta variables - these work
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    47
as placeholders. *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    48
fun dest_fake_fix_name n = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    49
    case (explode n) of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    50
      ("@" :: realchars) => implode realchars
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    51
    | _ => n;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
fun is_fake_fix_name n = (hd (explode n) = "@");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    53
fun mk_fake_fix_name n = "@" ^ n;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
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
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
(* fake free variable names for meta level bound variables *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    58
fun dest_fake_all_name n = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    59
    case (explode n) of 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    60
      ("+" :: realchars) => implode realchars
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
    | _ => n;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
fun is_fake_all_name n = (hd (explode n) = "+");
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
fun mk_fake_all_name n = "+" ^ n;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    64
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
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
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68
(* Ys and Ts not used, Ns are real names of faked local bounds, the
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    69
idea is that this will be mapped to free variables thus if a free
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    70
variable is a faked local bound then we change it to being a meta
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    71
variable so that it can later be instantiated *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    72
(* FIXME: rename this - avoid the word fix! *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    73
(* note we are not really "fix"'ing the free, more like making it variable! *)
15814
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    74
(* fun trymkvar_of_fakefree (Ns, Ts) Ys (n,ty) = 
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    75
    if n mem Ns then Var((n,0),ty) else Free (n,ty);
d65f461c8672 lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents: 15570
diff changeset
    76
*)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    77
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    78
(* make a var into a fixed free (ie prefixed with "@") *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    79
fun mk_fakefixvar Ts ((n,i),ty) = Free(mk_fake_fix_name n, ty);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    80
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
(* mk_frees_bound: string list -> Term.term -> Term.term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    83
(* This function changes free variables to being represented as bound
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    84
variables if the free's variable name is in the given list. The debruijn 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    85
index is simply the position in the list *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    86
(* THINKABOUT: danger of an existing free variable with the same name: fix
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    87
this so that name conflict are avoided automatically! In the meantime,
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    88
don't have free variables named starting with a ":" *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    89
fun bounds_of_fakefrees Ys (a $ b) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    90
    (bounds_of_fakefrees Ys a) $ (bounds_of_fakefrees Ys b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    91
  | bounds_of_fakefrees Ys (Abs(n,ty,t)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    92
    Abs(n,ty, bounds_of_fakefrees (n::Ys) t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    93
  | bounds_of_fakefrees Ys (Free (n,ty)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    94
    let fun try_mk_bound_of_free (i,[]) = Free (n,ty)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    95
          | try_mk_bound_of_free (i,(y::ys)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    96
            if n = y then Bound i else try_mk_bound_of_free (i+1,ys)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    97
    in try_mk_bound_of_free (0,Ys) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    98
  | bounds_of_fakefrees Ys t = t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    99
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   100
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   101
(* map a function f onto each free variables *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   102
fun map_to_frees f Ys (a $ b) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   103
    (map_to_frees f Ys a) $ (map_to_frees f Ys b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   104
  | map_to_frees f Ys (Abs(n,ty,t)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   105
    Abs(n,ty, map_to_frees f ((n,ty)::Ys) t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   106
  | map_to_frees f Ys (Free a) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   107
    f Ys a
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   108
  | map_to_frees f Ys t = t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   109
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   110
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   111
(* map a function f onto each meta variable  *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   112
fun map_to_vars f Ys (a $ b) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   113
    (map_to_vars f Ys a) $ (map_to_vars f Ys b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   114
  | map_to_vars f Ys (Abs(n,ty,t)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   115
    Abs(n,ty, map_to_vars f ((n,ty)::Ys) t)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   116
  | map_to_vars f Ys (Var a) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   117
    f Ys a
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   118
  | map_to_vars f Ys t = t;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   119
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   120
(* map a function f onto each free variables *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   121
fun map_to_alls f (Const("all",allty) $ Abs(n,ty,t)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   122
    let val (n2,ty2) = f (n,ty) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   123
    in (Const("all",allty) $ Abs(n2,ty2,map_to_alls f t)) end
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   124
  | map_to_alls f x = x;
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
(* map a function f to each type variable in a term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   127
(* implicit arg: term *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   128
fun map_to_term_tvars f =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   129
    Term.map_term_types (fn TVar(ix,ty) => f (ix,ty) | x => x);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   130
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   131
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   132
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   133
(* what if a param desn't occur in the concl? think about! Note: This
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   134
simply fixes meta level univ bound vars as Frees.  At the end, we will
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   135
change them back to schematic vars that will then unify
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   136
appropriactely, ie with unfake_vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   137
fun fake_concl_of_goal gt i = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   138
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   139
      val prems = Logic.strip_imp_prems gt
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   140
      val sgt = List.nth (prems, i - 1)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   141
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   142
      val tbody = Logic.strip_imp_concl (Term.strip_all_body sgt)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   143
      val tparams = Term.strip_all_vars sgt
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   144
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   145
      val fakefrees = map (fn (n, ty) => Free(mk_fake_all_name n, ty)) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   146
                          tparams
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   147
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   148
      Term.subst_bounds (rev fakefrees,tbody)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   149
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   150
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   151
(* what if a param desn't occur in the concl? think about! Note: This
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   152
simply fixes meta level univ bound vars as Frees.  At the end, we will
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   153
change them back to schematic vars that will then unify
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   154
appropriactely, ie with unfake_vars *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   155
fun fake_goal gt i = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   156
    let 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   157
      val prems = Logic.strip_imp_prems gt
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   158
      val sgt = List.nth (prems, i - 1)
15481
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
      val tbody = Term.strip_all_body sgt
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   161
      val tparams = Term.strip_all_vars sgt
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   162
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   163
      val fakefrees = map (fn (n, ty) => Free(mk_fake_all_name n, ty)) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   164
                          tparams
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   165
    in
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   166
      Term.subst_bounds (rev fakefrees,tbody)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   167
    end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   168
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   169
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   170
(* hand written - for some reason the Isabelle version in drule is broken!
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   171
Example? something to do with Bin Yangs examples? 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   172
 *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   173
fun rename_term_bvars ns (Abs(s,ty,t)) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   174
    let val s2opt = Library.find_first (fn (x,y) => s = x) ns
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   175
    in case s2opt of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   176
         NONE => (Abs(s,ty,rename_term_bvars  ns t))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   177
       | SOME (_,s2) => Abs(s2,ty, rename_term_bvars ns t) end
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   178
  | rename_term_bvars ns (a$b) = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   179
    (rename_term_bvars ns a) $ (rename_term_bvars ns b)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   180
  | rename_term_bvars _ x = x;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   181
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   182
fun rename_thm_bvars ns th = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   183
    let val t = Thm.prop_of th 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   184
    in Thm.rename_boundvars t (rename_term_bvars ns t) th end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   185
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   186
(* Finish this to show how it breaks! (raises the exception): 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   187
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   188
exception rename_thm_bvars_exp of ((string * string) list * Thm.thm)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   189
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   190
    Drule.rename_bvars ns th 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   191
    handle TERM _ => raise rename_thm_bvars_exp (ns, th); 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   192
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   193
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
   194
end;