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