src/HOL/Nominal/nominal_fresh_fun.ML
author haftmann
Fri, 07 Dec 2007 15:07:59 +0100
changeset 25571 c9e39eafc7a0
parent 24558 419f7cde7f59
child 26337 44473c957672
permissions -rw-r--r--
instantiation target rather than legacy instance
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
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
     5
Provides a tactic to generate fresh names and
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
     6
a tactic to analyse instances of the fresh_fun.
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
     7
22729
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
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    10
(* First some functions that should be in the library *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    11
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    12
(* A tactical which applies a list of int -> tactic to the          *) 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    13
(* corresponding subgoals present after the application of          *) 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    14
(* another tactic.                                                  *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    15
(*                                                                  *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    16
(*  T THENL [A,B,C] is equivalent to T THEN (C 3 THEN B 2 THEN A 1) *) 
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    17
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    18
infix 1 THENL
22792
2443ae6dac7d fix sml compilation
narboux
parents: 22787
diff changeset
    19
fun tac THENL tacs =
22753
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    20
 tac THEN
49d7818e6161 modify fresh_fun_simp to ease debugging
narboux
parents: 22729
diff changeset
    21
  (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
    22
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    23
(* A tactic which only succeeds when the argument *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    24
(* tactic solves completely the specified subgoal *)
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
    25
fun SOLVEI t = t THEN_ALL_NEW (fn i => no_tac);
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
    26
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
    27
(* A version of TRY for int -> tactic *)
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
    28
fun TRY' tac i =  TRY (tac i);
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
    29
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    30
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
    31
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    32
    val thy = theory_of_thm st;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    33
    val cgoal = nth (cprems_of st) (i - 1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    34
    val {maxidx, ...} = rep_cterm cgoal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    35
    val j = maxidx + 1;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    36
    val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    37
    val ps = Logic.strip_params (term_of cgoal);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    38
    val Ts = map snd ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    39
    val tinst' = map (fn (t, u) =>
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    40
      (head_of (Logic.incr_indexes (Ts, j) t),
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    41
       list_abs (ps, u))) tinst;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    42
    val th' = instf
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    43
      (map (pairself (ctyp_of thy)) tyinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    44
      (map (pairself (cterm_of thy)) tinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    45
      (Thm.lift_rule cgoal th)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    46
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    47
    compose_tac (elim, th', nprems_of th) i st
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    48
  end handle Subscript => Seq.empty;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    49
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    50
val res_inst_tac_term = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    51
  gen_res_inst_tac_term (curry Thm.instantiate);
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
val res_inst_tac_term' = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    54
  gen_res_inst_tac_term (K Drule.cterm_instantiate) [];
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
fun cut_inst_tac_term' tinst th =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    57
  res_inst_tac_term' tinst false (Tactic.make_elim_preserve th);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    58
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    59
fun get_dyn_thm thy name atom_name = (PureThy.get_thm thy (Name name)) handle _ => 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    60
                                          raise ERROR ("The atom type "^atom_name^" is not defined."); 
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    61
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    62
(* End of function waiting to be in the library :o) *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    63
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    64
(* The theorems needed that are known at compile time. *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    65
val at_exists_fresh' = @{thm "at_exists_fresh'"};
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    66
val fresh_fun_app'   = @{thm "fresh_fun_app'"};
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    67
val fresh_prod       = @{thm "fresh_prod"};
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    68
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    69
(* A tactic to generate a name fresh for  all the free *) 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    70
(* variables and parameters of the goal                *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    71
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    72
fun generate_fresh_tac atom_name i thm =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    73
 let 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    74
   val thy = theory_of_thm thm;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    75
(* 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
    76
   val atom_basename = Sign.base_name atom_name;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    77
   val goal = List.nth(prems_of thm, i-1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    78
   val ps = Logic.strip_params goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    79
   val Ts = rev (map snd ps);
24271
499608101177 avoid low-level tsig;
wenzelm
parents: 23368
diff changeset
    80
   fun is_of_fs_name T = Sign.of_sort thy (T, Sign.intern_sort thy ["fs_"^atom_basename]); 
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    81
(* rebuild de bruijn indices *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    82
   val bvs = map_index (Bound o fst) ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    83
(* select variables of the right class *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    84
   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
    85
     (term_frees goal @ bvs);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    86
(* build the tuple *)
23368
ad690b9bca1c generate_fresh works even if there is no free variable in the goal
narboux
parents: 23094
diff changeset
    87
   val s = (Library.foldr1 (fn (v, s) =>
ad690b9bca1c generate_fresh works even if there is no free variable in the goal
narboux
parents: 23094
diff changeset
    88
     HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    89
   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
    90
   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
    91
   val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    92
(* find the variable we want to instantiate *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    93
   val x = hd (term_vars (prop_of exists_fresh'));
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    94
 in 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    95
   (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    96
   rtac fs_name_thm 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    97
   etac exE 1) thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    98
  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
    99
  end;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   100
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   101
fun get_inner_fresh_fun (Bound j) = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   102
  | get_inner_fresh_fun (v as Free _) = NONE 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   103
  | get_inner_fresh_fun (v as Var _)  = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   104
  | get_inner_fresh_fun (Const _) = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   105
  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t 
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   106
  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u) 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   107
                           = SOME T 
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   108
  | get_inner_fresh_fun (t $ u) = 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   109
     let val a = get_inner_fresh_fun u in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   110
     if a = NONE then get_inner_fresh_fun t else a 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   111
     end;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   112
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   113
(* This tactic generates a fresh name of the atom type *) 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   114
(* given by the innermost fresh_fun                    *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   115
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   116
fun generate_fresh_fun_tac i thm =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   117
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   118
    val goal = List.nth(prems_of thm, i-1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   119
    val atom_name_opt = get_inner_fresh_fun goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   120
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   121
  case atom_name_opt of 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   122
    NONE => all_tac thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   123
  | SOME atom_name  => generate_fresh_tac atom_name i thm               
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   124
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   125
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   126
(* Two substitution tactics which looks for the innermost occurence in 
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   127
   one assumption or in the conclusion *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   128
23065
ab28e8398670 search bottom up to get the inner fresh fun
narboux
parents: 23054
diff changeset
   129
val search_fun     = curry (Seq.flat o (uncurry EqSubst.searchf_bt_unify_valid));
ab28e8398670 search bottom up to get the inner fresh fun
narboux
parents: 23054
diff changeset
   130
val search_fun_asm = EqSubst.skip_first_asm_occs_search EqSubst.searchf_bt_unify_valid;
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   131
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   132
fun subst_inner_tac           ctx = EqSubst.eqsubst_tac' ctx search_fun;
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   133
fun subst_inner_asm_tac_aux i ctx = EqSubst.eqsubst_asm_tac' ctx search_fun_asm i;
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   134
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   135
(* A tactic to substitute in the first assumption 
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   136
   which contains an occurence. *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   137
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   138
fun subst_inner_asm_tac ctx th =  
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   139
   curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux 
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   140
                                                             (1 upto Thm.nprems_of th)))))) ctx th;
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   141
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   142
fun fresh_fun_tac no_asm i thm = 
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   143
  (* Find the variable we instantiate *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   144
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   145
    val thy = theory_of_thm thm;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   146
    val ctx = Context.init_proof thy;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   147
    val ss = simpset_of thy;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   148
    val abs_fresh = PureThy.get_thms thy (Name "abs_fresh");
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   149
    val fresh_perm_app = PureThy.get_thms thy (Name "fresh_perm_app");
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   150
    val ss' = ss addsimps fresh_prod::abs_fresh;
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   151
    val ss'' = ss' addsimps fresh_perm_app;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   152
    val x = hd (tl (term_vars (prop_of exI)));
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   153
    val goal = nth (prems_of thm) (i-1);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   154
    val atom_name_opt = get_inner_fresh_fun goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   155
    val n = List.length (Logic.strip_params goal);
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   156
    (* Here we rely on the fact that the variable introduced by generate_fresh_tac *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   157
    (* is the last one in the list, the inner one *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   158
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   159
  case atom_name_opt of 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   160
    NONE => all_tac thm
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   161
  | SOME atom_name  =>    
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   162
  let 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   163
    val atom_basename = Sign.base_name atom_name;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   164
    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
   165
    val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   166
    fun inst_fresh vars params i st =
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   167
   let val vars' = term_vars (prop_of st);
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   168
       val thy = theory_of_thm st;
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   169
   in case vars' \\ vars of 
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   170
     [x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st)
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   171
    | _ => error "fresh_fun_simp: Too many variables, please report."
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   172
  end
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   173
  in
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   174
  ((fn st =>
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   175
  let 
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   176
    val vars = term_vars (prop_of st);
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   177
    val params = Logic.strip_params (nth (prems_of st) (i-1))
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   178
    (* The tactics which solve the subgoals generated 
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   179
       by the conditionnal rewrite rule. *)
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   180
    val post_rewrite_tacs =  
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   181
          [rtac pt_name_inst,
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   182
           rtac at_name_inst,
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   183
           TRY' (SOLVEI (NominalPermeq.finite_guess_tac ss'')),
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   184
           inst_fresh vars params THEN'
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   185
           (TRY' (SOLVEI (NominalPermeq.fresh_guess_tac ss''))) THEN'
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   186
           (TRY' (SOLVEI (asm_full_simp_tac ss'')))] 
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   187
  in 
23094
f1df8d2da98a fix a bug : the semantics of no_asm was the opposite
narboux
parents: 23091
diff changeset
   188
   ((if no_asm then no_tac else
f1df8d2da98a fix a bug : the semantics of no_asm was the opposite
narboux
parents: 23091
diff changeset
   189
    (subst_inner_asm_tac ctx fresh_fun_app' i THENL post_rewrite_tacs)) 
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   190
    ORELSE
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   191
    (subst_inner_tac     ctx fresh_fun_app' i THENL post_rewrite_tacs)) st
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   192
  end)) thm
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   193
  
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   194
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   195
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   196
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   197
(* syntax for options, given "(no_asm)" will give back true, without
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   198
   gives back false *)
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   199
val options_syntax =
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   200
    (Args.parens (Args.$$$ "no_asm") >> (K true)) ||
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   201
     (Scan.succeed false);
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   202
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   203
val setup_generate_fresh =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   204
  Method.goal_args_ctxt' Args.tyname (fn ctxt => generate_fresh_tac) 
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   205
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   206
val setup_fresh_fun_simp =
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   207
  Method.simple_args options_syntax 
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   208
  (fn b => fn _ => Method.SIMPLE_METHOD (fresh_fun_tac b 1))