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