| author | wenzelm | 
| Mon, 13 Jun 2016 22:42:38 +0200 | |
| changeset 63297 | ce995deef4b0 | 
| parent 60792 | 722cb21ab680 | 
| child 69597 | ff784d5a5bfb | 
| 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 | ||
| 56230 | 8 | (* FIXME proper ML structure! *) | 
| 22729 | 9 | |
| 43278 | 10 | (* FIXME res_inst_tac mostly obsolete, cf. Subgoal.FOCUS *) | 
| 11 | ||
| 12 | (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) | |
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58950diff
changeset | 13 | fun gen_res_inst_tac_term ctxt instf tyinst tinst elim th i st = | 
| 22729 | 14 | let | 
| 15 | val cgoal = nth (cprems_of st) (i - 1); | |
| 59586 | 16 | val maxidx = Thm.maxidx_of_cterm cgoal; | 
| 22729 | 17 | val j = maxidx + 1; | 
| 18 | val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst; | |
| 59582 | 19 | val ps = Logic.strip_params (Thm.term_of cgoal); | 
| 22729 | 20 | val Ts = map snd ps; | 
| 21 | val tinst' = map (fn (t, u) => | |
| 59787 | 22 | (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 | 23 | fold_rev Term.abs ps u)) tinst; | 
| 22729 | 24 | val th' = instf | 
| 60792 | 25 | (map (apsnd (Thm.ctyp_of ctxt)) tyinst') | 
| 26 | (map (apsnd (Thm.cterm_of ctxt)) tinst') | |
| 22729 | 27 | (Thm.lift_rule cgoal th) | 
| 28 | in | |
| 59582 | 29 | compose_tac ctxt (elim, th', Thm.nprems_of th) i st | 
| 43278 | 30 | end handle General.Subscript => Seq.empty; | 
| 31 | (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) | |
| 22729 | 32 | |
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58950diff
changeset | 33 | fun res_inst_tac_term ctxt = | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60359diff
changeset | 34 | gen_res_inst_tac_term ctxt (fn instT => fn inst => | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60359diff
changeset | 35 | Thm.instantiate | 
| 60792 | 36 | (map (apfst dest_TVar) instT, | 
| 37 | map (apfst dest_Var) inst)); | |
| 22729 | 38 | |
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58950diff
changeset | 39 | fun res_inst_tac_term' ctxt = | 
| 60792 | 40 | gen_res_inst_tac_term ctxt | 
| 41 | (fn _ => fn inst => infer_instantiate ctxt (map (apfst (#1 o dest_Var)) inst)) []; | |
| 22729 | 42 | |
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58318diff
changeset | 43 | fun cut_inst_tac_term' ctxt tinst th = | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58950diff
changeset | 44 | res_inst_tac_term' ctxt tinst false (Rule_Insts.make_elim_preserve ctxt th); | 
| 22729 | 45 | |
| 26337 | 46 | 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 | 47 | Global_Theory.get_thm thy name handle ERROR _ => | 
| 27187 | 48 |     error ("The atom type "^atom_name^" is not defined.");
 | 
| 22729 | 49 | |
| 24558 | 50 | (* The theorems needed that are known at compile time. *) | 
| 51 | val at_exists_fresh' = @{thm "at_exists_fresh'"};
 | |
| 52 | val fresh_fun_app'   = @{thm "fresh_fun_app'"};
 | |
| 53 | val fresh_prod       = @{thm "fresh_prod"};
 | |
| 22729 | 54 | |
| 33034 | 55 | (* A tactic to generate a name fresh for all the free *) | 
| 24558 | 56 | (* variables and parameters of the goal *) | 
| 22729 | 57 | |
| 56230 | 58 | fun generate_fresh_tac ctxt atom_name = SUBGOAL (fn (goal, _) => | 
| 33034 | 59 | let | 
| 56230 | 60 | val thy = Proof_Context.theory_of ctxt; | 
| 22729 | 61 | (* 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 | 62 | val atom_basename = Long_Name.base_name atom_name; | 
| 22729 | 63 | val ps = Logic.strip_params goal; | 
| 64 | val Ts = rev (map snd ps); | |
| 35667 | 65 |    fun is_of_fs_name T = Sign.of_sort thy (T, [Sign.intern_class thy ("fs_"^atom_basename)]);
 | 
| 22729 | 66 | (* rebuild de bruijn indices *) | 
| 67 | val bvs = map_index (Bound o fst) ps; | |
| 68 | (* select variables of the right class *) | |
| 69 | val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t))) | |
| 44121 | 70 | (Misc_Legacy.term_frees goal @ bvs); | 
| 22729 | 71 | (* build the tuple *) | 
| 23368 
ad690b9bca1c
generate_fresh works even if there is no free variable in the goal
 narboux parents: 
23094diff
changeset | 72 | val s = (Library.foldr1 (fn (v, s) => | 
| 41519 
0940fff556a6
avoid catch-all exception handler, presumably TERM was meant here;
 wenzelm parents: 
39557diff
changeset | 73 | 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 | 74 | handle TERM _ => HOLogic.unit; | 
| 22729 | 75 |    val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename;
 | 
| 76 |    val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
 | |
| 77 | val exists_fresh' = at_name_inst_thm RS at_exists_fresh'; | |
| 78 | (* find the variable we want to instantiate *) | |
| 59582 | 79 | val x = hd (Misc_Legacy.term_vars (Thm.prop_of exists_fresh')); | 
| 33034 | 80 | in | 
| 56230 | 81 | fn st => | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58318diff
changeset | 82 | (cut_inst_tac_term' ctxt [(x,s)] exists_fresh' 1 THEN | 
| 60754 | 83 | resolve_tac ctxt [fs_name_thm] 1 THEN | 
| 84 | eresolve_tac ctxt [exE] 1) st | |
| 56230 | 85 | handle List.Empty => all_tac st (* if we collected no variables then we do nothing *) | 
| 86 | end) 1; | |
| 22729 | 87 | |
| 88 | fun get_inner_fresh_fun (Bound j) = NONE | |
| 33034 | 89 | | get_inner_fresh_fun (v as Free _) = NONE | 
| 22729 | 90 | | get_inner_fresh_fun (v as Var _) = NONE | 
| 91 | | get_inner_fresh_fun (Const _) = NONE | |
| 33034 | 92 | | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t | 
| 56253 | 93 |   | get_inner_fresh_fun (Const (@{const_name Nominal.fresh_fun},
 | 
| 94 |       Type(@{type_name fun},[Type (@{type_name fun},[Type (T,_),_]),_])) $ u) = SOME T
 | |
| 33034 | 95 | | get_inner_fresh_fun (t $ u) = | 
| 22729 | 96 | let val a = get_inner_fresh_fun u in | 
| 33034 | 97 | if a = NONE then get_inner_fresh_fun t else a | 
| 22729 | 98 | end; | 
| 99 | ||
| 33034 | 100 | (* This tactic generates a fresh name of the atom type *) | 
| 24558 | 101 | (* given by the innermost fresh_fun *) | 
| 22729 | 102 | |
| 56230 | 103 | fun generate_fresh_fun_tac ctxt = SUBGOAL (fn (goal, _) => | 
| 22729 | 104 | let | 
| 105 | val atom_name_opt = get_inner_fresh_fun goal; | |
| 106 | in | |
| 33034 | 107 | case atom_name_opt of | 
| 56230 | 108 | NONE => all_tac | 
| 109 | | SOME atom_name => generate_fresh_tac ctxt atom_name | |
| 110 | end) 1; | |
| 22729 | 111 | |
| 58318 | 112 | (* Two substitution tactics which looks for the innermost occurrence in | 
| 23091 
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
 narboux parents: 
23065diff
changeset | 113 | one assumption or in the conclusion *) | 
| 22729 | 114 | |
| 33034 | 115 | val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid); | 
| 23065 | 116 | 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 | 117 | |
| 33034 | 118 | fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun; | 
| 119 | 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 | 120 | |
| 33034 | 121 | (* A tactic to substitute in the first assumption | 
| 58318 | 122 | which contains an occurrence. *) | 
| 22729 | 123 | |
| 33034 | 124 | fun subst_inner_asm_tac ctxt th = | 
| 125 | curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux | |
| 126 | (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 | 127 | |
| 56230 | 128 | fun fresh_fun_tac ctxt no_asm = SUBGOAL (fn (goal, i) => | 
| 22729 | 129 | (* Find the variable we instantiate *) | 
| 130 | let | |
| 56230 | 131 | val thy = Proof_Context.theory_of ctxt; | 
| 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 | 132 | 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 | 133 | 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: 
51669diff
changeset | 134 | val simp_ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51669diff
changeset | 135 | ctxt addsimps (fresh_prod :: abs_fresh) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51669diff
changeset | 136 | addsimps fresh_perm_app; | 
| 59582 | 137 | val x = hd (tl (Misc_Legacy.term_vars (Thm.prop_of exI))); | 
| 22729 | 138 | val atom_name_opt = get_inner_fresh_fun goal; | 
| 33034 | 139 | val n = length (Logic.strip_params goal); | 
| 24558 | 140 | (* Here we rely on the fact that the variable introduced by generate_fresh_tac *) | 
| 141 | (* is the last one in the list, the inner one *) | |
| 22729 | 142 | in | 
| 33034 | 143 | case atom_name_opt of | 
| 56230 | 144 | NONE => all_tac | 
| 33034 | 145 | | SOME atom_name => | 
| 146 | let | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 147 | val atom_basename = Long_Name.base_name atom_name; | 
| 22729 | 148 |     val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename;
 | 
| 149 |     val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename;
 | |
| 22787 | 150 | fun inst_fresh vars params i st = | 
| 59582 | 151 | let val vars' = Misc_Legacy.term_vars (Thm.prop_of st); | 
| 33040 | 152 | in case subtract (op =) vars vars' of | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60359diff
changeset | 153 | [Var v] => | 
| 60359 | 154 | Seq.single | 
| 155 | (Thm.instantiate ([], | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60359diff
changeset | 156 | [(v, Thm.cterm_of ctxt (fold_rev Term.abs params (Bound 0)))]) st) | 
| 22787 | 157 | | _ => 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 | 158 | end | 
| 22729 | 159 | in | 
| 23091 
c91530f18d9c
add an option in fresh_fun_simp to prevent rewriting in assumptions
 narboux parents: 
23065diff
changeset | 160 | ((fn st => | 
| 33034 | 161 | let | 
| 59582 | 162 | val vars = Misc_Legacy.term_vars (Thm.prop_of st); | 
| 163 | val params = Logic.strip_params (nth (Thm.prems_of st) (i-1)) | |
| 33034 | 164 | (* 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 | 165 | by the conditionnal rewrite rule. *) | 
| 33034 | 166 | val post_rewrite_tacs = | 
| 60754 | 167 | [resolve_tac ctxt [pt_name_inst], | 
| 168 | resolve_tac ctxt [at_name_inst], | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51669diff
changeset | 169 | 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: 
22792diff
changeset | 170 | inst_fresh vars params THEN' | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51669diff
changeset | 171 | (TRY o SOLVED' (NominalPermeq.fresh_guess_tac simp_ctxt)) THEN' | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51669diff
changeset | 172 | (TRY o SOLVED' (asm_full_simp_tac simp_ctxt))] | 
| 33034 | 173 | in | 
| 23094 
f1df8d2da98a
fix a bug : the semantics of no_asm was the opposite
 narboux parents: 
23091diff
changeset | 174 | ((if no_asm then no_tac else | 
| 33034 | 175 | (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 | 176 | ORELSE | 
| 33034 | 177 | (subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st | 
| 56230 | 178 | end)) | 
| 22729 | 179 | end | 
| 56230 | 180 | end) | 
| 22729 | 181 |