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