src/HOL/Nominal/nominal_fresh_fun.ML
author wenzelm
Thu, 18 Apr 2013 17:07:01 +0200
changeset 51717 9e7d1c139569
parent 51669 7fbc4fc400d8
child 55951 c07d184aebe9
permissions -rw-r--r--
simplifier uses proper Proof.context instead of historic type simpset;
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
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28397
diff changeset
     2
    Authors:    Stefan Berghofer and Julien Narboux, TU Muenchen
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     3
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
     4
Provides a tactic to generate fresh names and
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
     5
a tactic to analyse instances of the fresh_fun.
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     6
*)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     7
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
     8
(* FIXME proper ML structure *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
     9
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
    10
(* FIXME res_inst_tac mostly obsolete, cf. Subgoal.FOCUS *)
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
    11
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
    12
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    13
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
    14
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    15
    val thy = theory_of_thm st;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    16
    val cgoal = nth (cprems_of st) (i - 1);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    17
    val {maxidx, ...} = rep_cterm cgoal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    18
    val j = maxidx + 1;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    19
    val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    20
    val ps = Logic.strip_params (term_of cgoal);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    21
    val Ts = map snd ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    22
    val tinst' = map (fn (t, u) =>
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    23
      (head_of (Logic.incr_indexes (Ts, j) t),
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44121
diff changeset
    24
       fold_rev Term.abs ps u)) tinst;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    25
    val th' = instf
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    26
      (map (pairself (ctyp_of thy)) tyinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    27
      (map (pairself (cterm_of thy)) tinst')
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    28
      (Thm.lift_rule cgoal th)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    29
  in
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    30
    compose_tac (elim, th', nprems_of th) i st
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
    31
  end handle General.Subscript => Seq.empty;
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42806
diff changeset
    32
(* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    33
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    34
val res_inst_tac_term =
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    35
  gen_res_inst_tac_term (curry Thm.instantiate);
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    36
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    37
val res_inst_tac_term' =
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    38
  gen_res_inst_tac_term (K Drule.cterm_instantiate) [];
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    39
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    40
fun cut_inst_tac_term' tinst th =
42806
4b660cdab9b7 modernized structure Rule_Insts;
wenzelm
parents: 42364
diff changeset
    41
  res_inst_tac_term' tinst false (Rule_Insts.make_elim_preserve th);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    42
26337
44473c957672 auxiliary dynamic_thm(s) for fact lookup;
wenzelm
parents: 24558
diff changeset
    43
fun get_dyn_thm thy name atom_name =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 36610
diff changeset
    44
  Global_Theory.get_thm thy name handle ERROR _ =>
27187
17b63e145986 use regular error function;
wenzelm
parents: 26343
diff changeset
    45
    error ("The atom type "^atom_name^" is not defined.");
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    46
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    47
(* The theorems needed that are known at compile time. *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    48
val at_exists_fresh' = @{thm "at_exists_fresh'"};
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    49
val fresh_fun_app'   = @{thm "fresh_fun_app'"};
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    50
val fresh_prod       = @{thm "fresh_prod"};
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    51
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    52
(* A tactic to generate a name fresh for  all the free *)
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    53
(* variables and parameters of the goal                *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    54
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    55
fun generate_fresh_tac atom_name i thm =
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    56
 let
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    57
   val thy = theory_of_thm thm;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    58
(* the parsing function returns a qualified name, we get back the base name *)
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
    59
   val atom_basename = Long_Name.base_name atom_name;
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
    60
   val goal = nth (prems_of thm) (i - 1);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    61
   val ps = Logic.strip_params goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    62
   val Ts = rev (map snd ps);
35667
aa59452c913c tuned -- eliminated Sign.intern_sort;
wenzelm
parents: 35360
diff changeset
    63
   fun is_of_fs_name T = Sign.of_sort thy (T, [Sign.intern_class thy ("fs_"^atom_basename)]);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    64
(* rebuild de bruijn indices *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    65
   val bvs = map_index (Bound o fst) ps;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    66
(* select variables of the right class *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    67
   val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t)))
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43278
diff changeset
    68
     (Misc_Legacy.term_frees goal @ bvs);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    69
(* build the tuple *)
23368
ad690b9bca1c generate_fresh works even if there is no free variable in the goal
narboux
parents: 23094
diff changeset
    70
   val s = (Library.foldr1 (fn (v, s) =>
41519
0940fff556a6 avoid catch-all exception handler, presumably TERM was meant here;
wenzelm
parents: 39557
diff changeset
    71
       HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs)
0940fff556a6 avoid catch-all exception handler, presumably TERM was meant here;
wenzelm
parents: 39557
diff changeset
    72
     handle TERM _ => HOLogic.unit;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    73
   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
    74
   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
    75
   val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    76
(* find the variable we want to instantiate *)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43278
diff changeset
    77
   val x = hd (Misc_Legacy.term_vars (prop_of exists_fresh'));
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    78
 in
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    79
   (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    80
   rtac fs_name_thm 1 THEN
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    81
   etac exE 1) thm
47060
e2741ec9ae36 prefer explicitly qualified exception List.Empty;
wenzelm
parents: 46219
diff changeset
    82
  handle List.Empty  => all_tac thm (* if we collected no variables then we do nothing *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    83
  end;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    84
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    85
fun get_inner_fresh_fun (Bound j) = NONE
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    86
  | get_inner_fresh_fun (v as Free _) = NONE
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    87
  | get_inner_fresh_fun (v as Var _)  = NONE
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    88
  | get_inner_fresh_fun (Const _) = NONE
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    89
  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    90
  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u)
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    91
                           = SOME T
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    92
  | get_inner_fresh_fun (t $ u) =
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    93
     let val a = get_inner_fresh_fun u in
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    94
     if a = NONE then get_inner_fresh_fun t else a
22729
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
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
    97
(* This tactic generates a fresh name of the atom type *)
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
    98
(* given by the innermost fresh_fun                    *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
    99
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   100
fun generate_fresh_fun_tac i thm =
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   101
  let
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
   102
    val goal = nth (prems_of thm) (i - 1);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   103
    val atom_name_opt = get_inner_fresh_fun goal;
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   104
  in
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   105
  case atom_name_opt of
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   106
    NONE => all_tac thm
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   107
  | SOME atom_name  => generate_fresh_tac atom_name i thm
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   108
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   109
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   110
(* 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
   111
   one assumption or in the conclusion *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   112
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   113
val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid);
23065
ab28e8398670 search bottom up to get the inner fresh fun
narboux
parents: 23054
diff changeset
   114
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
   115
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   116
fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun;
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   117
fun subst_inner_asm_tac_aux i ctxt = EqSubst.eqsubst_asm_tac' ctxt 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
   118
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   119
(* A tactic to substitute in the first assumption
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   120
   which contains an occurence. *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   121
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   122
fun subst_inner_asm_tac ctxt th =
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   123
  curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   124
            (1 upto Thm.nprems_of th)))))) ctxt th;
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   125
51669
7fbc4fc400d8 proper proof context;
wenzelm
parents: 47060
diff changeset
   126
fun fresh_fun_tac ctxt no_asm i thm =
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   127
  (* Find the variable we instantiate *)
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   128
  let
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   129
    val thy = theory_of_thm thm;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 36610
diff changeset
   130
    val abs_fresh = Global_Theory.get_thms thy "abs_fresh";
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 36610
diff changeset
   131
    val fresh_perm_app = Global_Theory.get_thms thy "fresh_perm_app";
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   132
    val simp_ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   133
      ctxt addsimps (fresh_prod :: abs_fresh)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   134
      addsimps fresh_perm_app;
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43278
diff changeset
   135
    val x = hd (tl (Misc_Legacy.term_vars (prop_of exI)));
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   136
    val goal = nth (prems_of thm) (i-1);
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   137
    val atom_name_opt = get_inner_fresh_fun goal;
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   138
    val n = length (Logic.strip_params goal);
24558
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   139
    (* Here we rely on the fact that the variable introduced by generate_fresh_tac *)
419f7cde7f59 some cleaning up
urbanc
parents: 24271
diff changeset
   140
    (* is the last one in the list, the inner one *)
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   141
  in
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   142
  case atom_name_opt of
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   143
    NONE => all_tac thm
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   144
  | SOME atom_name =>
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   145
  let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   146
    val atom_basename = Long_Name.base_name atom_name;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   147
    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
   148
    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
   149
    fun inst_fresh vars params i st =
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43278
diff changeset
   150
   let val vars' = Misc_Legacy.term_vars (prop_of st);
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   151
       val thy = theory_of_thm st;
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33034
diff changeset
   152
   in case subtract (op =) vars vars' of
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44121
diff changeset
   153
     [x] =>
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44121
diff changeset
   154
      Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (fold_rev Term.abs params (Bound 0)))]) st)
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   155
    | _ => error "fresh_fun_simp: Too many variables, please report."
46219
426ed18eba43 discontinued old-style Term.list_abs in favour of plain Term.abs;
wenzelm
parents: 44121
diff changeset
   156
   end
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   157
  in
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   158
  ((fn st =>
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   159
  let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43278
diff changeset
   160
    val vars = Misc_Legacy.term_vars (prop_of st);
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   161
    val params = Logic.strip_params (nth (prems_of st) (i-1))
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   162
    (* The tactics which solve the subgoals generated
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   163
       by the conditionnal rewrite rule. *)
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   164
    val post_rewrite_tacs =
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   165
          [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
   166
           rtac at_name_inst,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   167
           TRY o SOLVED' (NominalPermeq.finite_guess_tac simp_ctxt),
23054
d1bbe5afa279 change fresh_fun_simp to treat occurences in assumptions and try to solve the generated subgoals
narboux
parents: 22792
diff changeset
   168
           inst_fresh vars params THEN'
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   169
           (TRY o SOLVED' (NominalPermeq.fresh_guess_tac simp_ctxt)) THEN'
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51669
diff changeset
   170
           (TRY o SOLVED' (asm_full_simp_tac simp_ctxt))]
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   171
  in
23094
f1df8d2da98a fix a bug : the semantics of no_asm was the opposite
narboux
parents: 23091
diff changeset
   172
   ((if no_asm then no_tac else
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   173
    (subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i)))
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   174
    ORELSE
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   175
    (subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st
22787
85e7e32e6004 update fresh_fun_simp for debugging purposes
narboux
parents: 22774
diff changeset
   176
  end)) thm
33034
66ef64a5f122 eliminated THENL -- use THEN RANGE;
wenzelm
parents: 32149
diff changeset
   177
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   178
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   179
  end
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   180
23091
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   181
(* 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
   182
   gives back false *)
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   183
val options_syntax =
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   184
    (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
   185
     (Scan.succeed false);
c91530f18d9c add an option in fresh_fun_simp to prevent rewriting in assumptions
narboux
parents: 23065
diff changeset
   186
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   187
fun setup_generate_fresh x =
35360
df2b2168e43a clarified ProofContext.read_type_name/Args.type_name wrt strict logical constructors;
wenzelm
parents: 34885
diff changeset
   188
  (Args.goal_spec -- Args.type_name true >>
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   189
    (fn (quant, s) => K (SIMPLE_METHOD'' quant (generate_fresh_tac s)))) x;
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents:
diff changeset
   190
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   191
fun setup_fresh_fun_simp x =
51669
7fbc4fc400d8 proper proof context;
wenzelm
parents: 47060
diff changeset
   192
  (Scan.lift options_syntax >> (fn b => fn ctxt => SIMPLE_METHOD' (fresh_fun_tac ctxt b))) x;
30549
d2d7874648bd simplified method setup;
wenzelm
parents: 30510
diff changeset
   193