src/HOL/Nominal/nominal_fresh_fun.ML
author wenzelm
Mon, 03 May 2010 14:25:56 +0200
changeset 36633 bafd82950e24
parent 35674 aa59452c913c
child 39814 fe5722fce758
permissions -rw-r--r--
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
narboux@22729
     1
(*  Title:      HOL/Nominal/nominal_fresh_fun.ML
wenzelm@29265
     2
    Authors:    Stefan Berghofer and Julien Narboux, TU Muenchen
narboux@22729
     3
urbanc@24558
     4
Provides a tactic to generate fresh names and
urbanc@24558
     5
a tactic to analyse instances of the fresh_fun.
narboux@22729
     6
*)
narboux@22729
     7
urbanc@24558
     8
(* First some functions that should be in the library *)
narboux@22729
     9
narboux@22729
    10
fun gen_res_inst_tac_term instf tyinst tinst elim th i st =
narboux@22729
    11
  let
narboux@22729
    12
    val thy = theory_of_thm st;
narboux@22729
    13
    val cgoal = nth (cprems_of st) (i - 1);
narboux@22729
    14
    val {maxidx, ...} = rep_cterm cgoal;
narboux@22729
    15
    val j = maxidx + 1;
narboux@22729
    16
    val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst;
narboux@22729
    17
    val ps = Logic.strip_params (term_of cgoal);
narboux@22729
    18
    val Ts = map snd ps;
narboux@22729
    19
    val tinst' = map (fn (t, u) =>
narboux@22729
    20
      (head_of (Logic.incr_indexes (Ts, j) t),
narboux@22729
    21
       list_abs (ps, u))) tinst;
narboux@22729
    22
    val th' = instf
narboux@22729
    23
      (map (pairself (ctyp_of thy)) tyinst')
narboux@22729
    24
      (map (pairself (cterm_of thy)) tinst')
narboux@22729
    25
      (Thm.lift_rule cgoal th)
narboux@22729
    26
  in
narboux@22729
    27
    compose_tac (elim, th', nprems_of th) i st
narboux@22729
    28
  end handle Subscript => Seq.empty;
narboux@22729
    29
wenzelm@33034
    30
val res_inst_tac_term =
narboux@22729
    31
  gen_res_inst_tac_term (curry Thm.instantiate);
narboux@22729
    32
wenzelm@33034
    33
val res_inst_tac_term' =
narboux@22729
    34
  gen_res_inst_tac_term (K Drule.cterm_instantiate) [];
narboux@22729
    35
narboux@22729
    36
fun cut_inst_tac_term' tinst th =
wenzelm@27239
    37
  res_inst_tac_term' tinst false (RuleInsts.make_elim_preserve th);
narboux@22729
    38
wenzelm@26337
    39
fun get_dyn_thm thy name atom_name =
wenzelm@26343
    40
  PureThy.get_thm thy name handle ERROR _ =>
wenzelm@27187
    41
    error ("The atom type "^atom_name^" is not defined.");
narboux@22729
    42
urbanc@24558
    43
(* End of function waiting to be in the library :o) *)
narboux@22729
    44
urbanc@24558
    45
(* The theorems needed that are known at compile time. *)
urbanc@24558
    46
val at_exists_fresh' = @{thm "at_exists_fresh'"};
urbanc@24558
    47
val fresh_fun_app'   = @{thm "fresh_fun_app'"};
urbanc@24558
    48
val fresh_prod       = @{thm "fresh_prod"};
narboux@22729
    49
wenzelm@33034
    50
(* A tactic to generate a name fresh for  all the free *)
urbanc@24558
    51
(* variables and parameters of the goal                *)
narboux@22729
    52
narboux@22729
    53
fun generate_fresh_tac atom_name i thm =
wenzelm@33034
    54
 let
narboux@22729
    55
   val thy = theory_of_thm thm;
narboux@22729
    56
(* the parsing function returns a qualified name, we get back the base name *)
wenzelm@30364
    57
   val atom_basename = Long_Name.base_name atom_name;
narboux@22729
    58
   val goal = List.nth(prems_of thm, i-1);
narboux@22729
    59
   val ps = Logic.strip_params goal;
narboux@22729
    60
   val Ts = rev (map snd ps);
wenzelm@35674
    61
   fun is_of_fs_name T = Sign.of_sort thy (T, [Sign.intern_class thy ("fs_"^atom_basename)]);
narboux@22729
    62
(* rebuild de bruijn indices *)
narboux@22729
    63
   val bvs = map_index (Bound o fst) ps;
narboux@22729
    64
(* select variables of the right class *)
narboux@22729
    65
   val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t)))
wenzelm@29265
    66
     (OldTerm.term_frees goal @ bvs);
narboux@22729
    67
(* build the tuple *)
narboux@23368
    68
   val s = (Library.foldr1 (fn (v, s) =>
wenzelm@28397
    69
     HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle _ => HOLogic.unit ;  (* FIXME avoid handle _ *)
narboux@22729
    70
   val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename;
narboux@22729
    71
   val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
narboux@22729
    72
   val exists_fresh' = at_name_inst_thm RS at_exists_fresh';
narboux@22729
    73
(* find the variable we want to instantiate *)
wenzelm@29265
    74
   val x = hd (OldTerm.term_vars (prop_of exists_fresh'));
wenzelm@33034
    75
 in
narboux@22729
    76
   (cut_inst_tac_term' [(x,s)] exists_fresh' 1 THEN
narboux@22729
    77
   rtac fs_name_thm 1 THEN
narboux@22729
    78
   etac exE 1) thm
narboux@22729
    79
  handle Empty  => all_tac thm (* if we collected no variables then we do nothing *)
narboux@22729
    80
  end;
narboux@22729
    81
narboux@22729
    82
fun get_inner_fresh_fun (Bound j) = NONE
wenzelm@33034
    83
  | get_inner_fresh_fun (v as Free _) = NONE
narboux@22729
    84
  | get_inner_fresh_fun (v as Var _)  = NONE
narboux@22729
    85
  | get_inner_fresh_fun (Const _) = NONE
wenzelm@33034
    86
  | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t
wenzelm@33034
    87
  | get_inner_fresh_fun (Const ("Nominal.fresh_fun",Type("fun",[Type ("fun",[Type (T,_),_]),_])) $ u)
wenzelm@33034
    88
                           = SOME T
wenzelm@33034
    89
  | get_inner_fresh_fun (t $ u) =
narboux@22729
    90
     let val a = get_inner_fresh_fun u in
wenzelm@33034
    91
     if a = NONE then get_inner_fresh_fun t else a
narboux@22729
    92
     end;
narboux@22729
    93
wenzelm@33034
    94
(* This tactic generates a fresh name of the atom type *)
urbanc@24558
    95
(* given by the innermost fresh_fun                    *)
narboux@22729
    96
narboux@22729
    97
fun generate_fresh_fun_tac i thm =
narboux@22729
    98
  let
narboux@22729
    99
    val goal = List.nth(prems_of thm, i-1);
narboux@22729
   100
    val atom_name_opt = get_inner_fresh_fun goal;
narboux@22729
   101
  in
wenzelm@33034
   102
  case atom_name_opt of
narboux@22729
   103
    NONE => all_tac thm
wenzelm@33034
   104
  | SOME atom_name  => generate_fresh_tac atom_name i thm
narboux@22729
   105
  end
narboux@22729
   106
wenzelm@33034
   107
(* Two substitution tactics which looks for the innermost occurence in
narboux@23091
   108
   one assumption or in the conclusion *)
narboux@22729
   109
wenzelm@33034
   110
val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid);
narboux@23065
   111
val search_fun_asm = EqSubst.skip_first_asm_occs_search EqSubst.searchf_bt_unify_valid;
narboux@23054
   112
wenzelm@33034
   113
fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun;
wenzelm@33034
   114
fun subst_inner_asm_tac_aux i ctxt = EqSubst.eqsubst_asm_tac' ctxt search_fun_asm i;
narboux@23054
   115
wenzelm@33034
   116
(* A tactic to substitute in the first assumption
narboux@23091
   117
   which contains an occurence. *)
narboux@22729
   118
wenzelm@33034
   119
fun subst_inner_asm_tac ctxt th =
wenzelm@33034
   120
  curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux
wenzelm@33034
   121
            (1 upto Thm.nprems_of th)))))) ctxt th;
narboux@23091
   122
wenzelm@33034
   123
fun fresh_fun_tac no_asm i thm =
narboux@22729
   124
  (* Find the variable we instantiate *)
narboux@22729
   125
  let
narboux@22729
   126
    val thy = theory_of_thm thm;
wenzelm@36633
   127
    val ctxt = ProofContext.init_global thy;
wenzelm@32149
   128
    val ss = global_simpset_of thy;
wenzelm@26343
   129
    val abs_fresh = PureThy.get_thms thy "abs_fresh";
wenzelm@26343
   130
    val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app";
narboux@22729
   131
    val ss' = ss addsimps fresh_prod::abs_fresh;
narboux@23054
   132
    val ss'' = ss' addsimps fresh_perm_app;
wenzelm@29265
   133
    val x = hd (tl (OldTerm.term_vars (prop_of exI)));
narboux@22787
   134
    val goal = nth (prems_of thm) (i-1);
narboux@22729
   135
    val atom_name_opt = get_inner_fresh_fun goal;
wenzelm@33034
   136
    val n = length (Logic.strip_params goal);
urbanc@24558
   137
    (* Here we rely on the fact that the variable introduced by generate_fresh_tac *)
urbanc@24558
   138
    (* is the last one in the list, the inner one *)
narboux@22729
   139
  in
wenzelm@33034
   140
  case atom_name_opt of
narboux@22729
   141
    NONE => all_tac thm
wenzelm@33034
   142
  | SOME atom_name =>
wenzelm@33034
   143
  let
wenzelm@30364
   144
    val atom_basename = Long_Name.base_name atom_name;
narboux@22729
   145
    val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
narboux@22729
   146
    val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
narboux@22787
   147
    fun inst_fresh vars params i st =
wenzelm@29265
   148
   let val vars' = OldTerm.term_vars (prop_of st);
narboux@22787
   149
       val thy = theory_of_thm st;
haftmann@33040
   150
   in case subtract (op =) vars vars' of
narboux@22787
   151
     [x] => Seq.single (Thm.instantiate ([],[(cterm_of thy x,cterm_of thy (list_abs (params,Bound 0)))]) st)
narboux@22787
   152
    | _ => error "fresh_fun_simp: Too many variables, please report."
narboux@22787
   153
  end
narboux@22729
   154
  in
narboux@23091
   155
  ((fn st =>
wenzelm@33034
   156
  let
wenzelm@29265
   157
    val vars = OldTerm.term_vars (prop_of st);
narboux@22787
   158
    val params = Logic.strip_params (nth (prems_of st) (i-1))
wenzelm@33034
   159
    (* The tactics which solve the subgoals generated
narboux@23091
   160
       by the conditionnal rewrite rule. *)
wenzelm@33034
   161
    val post_rewrite_tacs =
narboux@23054
   162
          [rtac pt_name_inst,
narboux@23054
   163
           rtac at_name_inst,
wenzelm@34885
   164
           TRY o SOLVED' (NominalPermeq.finite_guess_tac ss''),
narboux@23054
   165
           inst_fresh vars params THEN'
wenzelm@34885
   166
           (TRY o SOLVED' (NominalPermeq.fresh_guess_tac ss'')) THEN'
wenzelm@34885
   167
           (TRY o SOLVED' (asm_full_simp_tac ss''))]
wenzelm@33034
   168
  in
narboux@23094
   169
   ((if no_asm then no_tac else
wenzelm@33034
   170
    (subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i)))
narboux@23091
   171
    ORELSE
wenzelm@33034
   172
    (subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st
narboux@22787
   173
  end)) thm
wenzelm@33034
   174
narboux@22729
   175
  end
narboux@22729
   176
  end
narboux@22729
   177
narboux@23091
   178
(* syntax for options, given "(no_asm)" will give back true, without
narboux@23091
   179
   gives back false *)
narboux@23091
   180
val options_syntax =
narboux@23091
   181
    (Args.parens (Args.$$$ "no_asm") >> (K true)) ||
narboux@23091
   182
     (Scan.succeed false);
narboux@23091
   183
wenzelm@30549
   184
fun setup_generate_fresh x =
wenzelm@35360
   185
  (Args.goal_spec -- Args.type_name true >>
wenzelm@30549
   186
    (fn (quant, s) => K (SIMPLE_METHOD'' quant (generate_fresh_tac s)))) x;
narboux@22729
   187
wenzelm@30549
   188
fun setup_fresh_fun_simp x =
wenzelm@30549
   189
  (Scan.lift options_syntax >> (fn b => K (SIMPLE_METHOD' (fresh_fun_tac b)))) x;
wenzelm@30549
   190