src/HOL/Nominal/nominal_fresh_fun.ML
author narboux
Tue, 24 Apr 2007 14:01:23 +0200
changeset 22774 8c64803fae48
parent 22753 49d7818e6161
child 22787 85e7e32e6004
permissions -rw-r--r--
adds op in front of an infix to fix SML compilation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_fresh_fun.ML
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     2
    ID:         $Id$
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     3
    Authors:    Stefan Berghofer, Julien Narboux, TU Muenchen
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     4
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     5
A tactic to generate fresh names.
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     6
A tactic to get rid of the fresh_fun.
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     7
*)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     8
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     9
(* First some functions that could be
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    10
 in the library *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    11
22774
8c64803fae48 adds op in front of an infix to fix SML compilation
narboux
parents: 22753
diff changeset
    12
(* A tactical which applies a list of int -> tactic to the corresponding subgoals present after the application of another tactic. 
8c64803fae48 adds op in front of an infix to fix SML compilation
narboux
parents: 22753
diff changeset
    13
T THENL [A,B,C] is equivalent to T THEN (C 3 THEN B 2 THEN A 1) 
8c64803fae48 adds op in front of an infix to fix SML compilation
narboux
parents: 22753
diff changeset
    14
*)
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    15
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    16
infix 1 THENL
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    17
22774
8c64803fae48 adds op in front of an infix to fix SML compilation
narboux
parents: 22753
diff changeset
    18
fun (op THENL) (tac,tacs) =
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    19
 tac THEN
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    20
  (EVERY (map  (fn (tac,i) => tac i) (rev tacs ~~ (length tacs downto 1))))
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    21
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    22
fun gen_res_inst_tac_term instf tyinst tinst elim th i st =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    23
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    24
    val thy = theory_of_thm st;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    25
    val cgoal = nth (cprems_of st) (i - 1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    26
    val {maxidx, ...} = rep_cterm cgoal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    27
    val j = maxidx + 1;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    28
    val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    29
    val ps = Logic.strip_params (term_of cgoal);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    30
    val Ts = map snd ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    31
    val tinst' = map (fn (t, u) =>
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    32
      (head_of (Logic.incr_indexes (Ts, j) t),
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    33
       list_abs (ps, u))) tinst;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    34
    val th' = instf
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    35
      (map (pairself (ctyp_of thy)) tyinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    36
      (map (pairself (cterm_of thy)) tinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    37
      (Thm.lift_rule cgoal th)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    38
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    39
    compose_tac (elim, th', nprems_of th) i st
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    40
  end handle Subscript => Seq.empty;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    41
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    42
val res_inst_tac_term = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    43
  gen_res_inst_tac_term (curry Thm.instantiate);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    44
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    45
val res_inst_tac_term' = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    46
  gen_res_inst_tac_term (K Drule.cterm_instantiate) [];
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    47
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    48
fun cut_inst_tac_term' tinst th =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    49
  res_inst_tac_term' tinst false (Tactic.make_elim_preserve th);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    50
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    51
fun get_dyn_thm thy name atom_name = (PureThy.get_thm thy (Name name)) handle _ => raise ERROR ("The atom type "^atom_name^" is not defined."); 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    52
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    53
(* End of function waiting to be
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    54
 in the library *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    55
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    56
(* The theorems needed that are known at
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    57
compile time. *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    58
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    59
val at_exists_fresh' = thm "at_exists_fresh'";
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    60
val fresh_fun_app = thm "fresh_fun_app";
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    61
val fresh_prod = thm "fresh_prod";
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    62
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    63
(* A tactic to generate a name fresh for 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    64
all the free variables and parameters of the goal *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    65
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    66
fun generate_fresh_tac atom_name i thm =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    67
 let 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    68
   val thy = theory_of_thm thm;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    69
(* the parsing function returns a qualified name, we get back the base name *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    70
   val atom_basename = Sign.base_name atom_name;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    71
   val goal = List.nth(prems_of thm, i-1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    72
   val ps = Logic.strip_params goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    73
   val Ts = rev (map snd ps);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    74
   fun is_of_fs_name T = Type.of_sort (Sign.tsig_of thy)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    75
     (T, Sign.intern_sort thy ["fs_"^atom_basename]); 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    76
(* rebuild de bruijn indices *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    77
   val bvs = map_index (Bound o fst) ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    78
(* select variables of the right class *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    79
   val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t)))
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    80
     (term_frees goal @ bvs);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    81
(* build the tuple *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    82
   val s = Library.foldr1 (fn (v, s) =>
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    83
     HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    84
       vs;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    85
   val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    86
   val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    87
   val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    88
(* find the variable we want to instantiate *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    89
   val x = hd (term_vars (prop_of exists_fresh'));
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    90
 in 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    91
   (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    92
   rtac fs_name_thm 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    93
   etac exE 1) thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    94
  handle Empty  => all_tac thm (* if we collected no variables then we do nothing *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    95
  end;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    96
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    97
fun get_inner_fresh_fun (Bound j) = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    98
  | get_inner_fresh_fun (v as Free _) = NONE 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    99
  | get_inner_fresh_fun (v as Var _)  = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   100
  | get_inner_fresh_fun (Const _) = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   101
  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   102
  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u) = SOME T 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   103
  | get_inner_fresh_fun (t $ u) = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   104
     let val a = get_inner_fresh_fun u in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   105
     if a = NONE then get_inner_fresh_fun t else a 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   106
     end;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   107
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   108
(* This tactic generates a fresh name 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   109
of the atom type given by the inner most fresh_fun *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   110
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   111
fun generate_fresh_fun_tac i thm =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   112
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   113
    val goal = List.nth(prems_of thm, i-1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   114
    val atom_name_opt = get_inner_fresh_fun goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   115
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   116
  case atom_name_opt of 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   117
    NONE => all_tac thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   118
  | SOME atom_name  => generate_fresh_tac atom_name i thm               
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   119
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   120
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   121
(* A substitution tactic, FIXME : change the search function to get the inner occurence of fresh_fun (bottom -> up search ) *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   122
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   123
fun subst_outer_tac ctx = EqSubst.eqsubst_tac' ctx (curry (Seq.flat o (uncurry EqSubst.searchf_lr_unify_valid)));
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   124
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   125
fun fresh_fun_tac i thm = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   126
  (* Find the variable we instantiate *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   127
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   128
    val thy = theory_of_thm thm;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   129
    val ctx = Context.init_proof thy;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   130
    val ss = simpset_of thy;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   131
    val abs_fresh = PureThy.get_thms thy (Name "abs_fresh");
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   132
    val ss' = ss addsimps fresh_prod::abs_fresh;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   133
    val x = hd (tl (term_vars (prop_of exI)));
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   134
    val goal = List.nth(prems_of thm, i-1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   135
    val atom_name_opt = get_inner_fresh_fun goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   136
    val n = List.length (Logic.strip_params goal);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   137
    (* Here we rely on the fact that the variable introduced by generate_fresh_tac is the last one in the list, the inner one *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   138
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   139
  case atom_name_opt of 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   140
    NONE => all_tac thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   141
  | SOME atom_name  =>    
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   142
  let 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   143
    val atom_basename = Sign.base_name atom_name;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   144
    val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   145
    val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   146
    fun tac_a i = 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   147
      res_inst_tac_term' [(x,Bound n)] false exI i THEN
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   148
      asm_full_simp_tac ss' i                      THEN 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   149
      NominalPermeq.fresh_guess_tac HOL_ss i;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   150
  in
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   151
  (generate_fresh_tac atom_name i      THEN 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   152
   subst_outer_tac ctx fresh_fun_app i THENL 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   153
   [rtac pt_name_inst,
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   154
    rtac at_name_inst, 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   155
    (fn i => TRY (NominalPermeq.finite_guess_tac HOL_ss i)),   
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   156
    tac_a, 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   157
    (fn i => TRY (NominalPermeq.fresh_guess_tac  HOL_ss i)),
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   158
    (fn i => auto_tac (HOL_cs, HOL_ss)) 
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
   159
]) thm 
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   160
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   161
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   162
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   163
val setup_generate_fresh =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   164
  Method.goal_args_ctxt' Args.tyname (fn ctxt => generate_fresh_tac) 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   165
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   166
val setup_fresh_fun_simp =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   167
  Method.no_args (Method.SIMPLE_METHOD (fresh_fun_tac 1))