| author | wenzelm | 
| Sat, 13 Nov 2021 17:26:35 +0100 | |
| changeset 74779 | 5fca489a6ac1 | 
| parent 74524 | 8ee3d5d3c1bf | 
| child 78095 | bc42c074e58f | 
| permissions | -rw-r--r-- | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 1 | (* Title: HOL/Nominal/nominal_inductive.ML | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 3 | |
| 22530 | 4 | Infrastructure for proving equivariance and strong induction theorems | 
| 5 | for inductive predicates involving nominal datatypes. | |
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 6 | *) | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 7 | |
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 8 | signature NOMINAL_INDUCTIVE = | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 9 | sig | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 10 | val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 11 | val prove_eqvt: string -> string list -> local_theory -> local_theory | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 12 | end | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 13 | |
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 14 | structure NominalInductive : NOMINAL_INDUCTIVE = | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 15 | struct | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 16 | |
| 59940 
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
 wenzelm parents: 
59936diff
changeset | 17 | val inductive_forall_def = @{thm HOL.induct_forall_def};
 | 
| 33772 
b6a1feca2ac2
use of thm-antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
33671diff
changeset | 18 | val inductive_atomize = @{thms induct_atomize};
 | 
| 
b6a1feca2ac2
use of thm-antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
33671diff
changeset | 19 | val inductive_rulify = @{thms induct_rulify};
 | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 20 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
39557diff
changeset | 21 | fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 22 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 23 | fun atomize_conv ctxt = | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
39557diff
changeset | 24 | Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 25 | (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 26 | fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); | 
| 24832 | 27 | fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 28 | (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)); | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 29 | |
| 44929 | 30 | fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 31 | |
| 39159 | 32 | val fresh_prod = @{thm fresh_prod};
 | 
| 22530 | 33 | |
| 44689 | 34 | val perm_bool = mk_meta_eq @{thm perm_bool_def};
 | 
| 39159 | 35 | val perm_boolI = @{thm perm_boolI};
 | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 36 | val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb | 
| 59582 | 37 | (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 38 | |
| 60787 | 39 | fun mk_perm_bool ctxt pi th = | 
| 40 | th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; | |
| 25824 | 41 | |
| 60359 | 42 | fun mk_perm_bool_simproc names = | 
| 69597 | 43 | Simplifier.make_simproc \<^context> "perm_bool" | 
| 44 |    {lhss = [\<^term>\<open>perm pi x\<close>],
 | |
| 61144 | 45 | proc = fn _ => fn _ => fn ct => | 
| 46 | (case Thm.term_of ct of | |
| 69597 | 47 | Const (\<^const_name>\<open>Nominal.perm\<close>, _) $ _ $ t => | 
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 48 | if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) | 
| 61144 | 49 | then SOME perm_bool else NONE | 
| 62913 | 50 | | _ => NONE)}; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 51 | |
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 52 | fun transp ([] :: _) = [] | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 53 | | transp xs = map hd xs :: transp (map tl xs); | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 54 | |
| 22530 | 55 | fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of | 
| 56 | (Const (s, T), ts) => (case strip_type T of | |
| 57 | (Ts, Type (tname, _)) => | |
| 31938 | 58 | (case NominalDatatype.get_nominal_datatype thy tname of | 
| 22530 | 59 | NONE => fold (add_binders thy i) ts bs | 
| 60 |            | SOME {descr, index, ...} => (case AList.lookup op =
 | |
| 61 | (#3 (the (AList.lookup op = descr index))) s of | |
| 62 | NONE => fold (add_binders thy i) ts bs | |
| 63 | | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => | |
| 64 | let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' | |
| 65 | in (add_binders thy i u | |
| 66 | (fold (fn (u, T) => | |
| 67 | if exists (fn j => j < i) (loose_bnos u) then I | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58112diff
changeset | 68 | else insert (op aconv o apply2 fst) | 
| 22530 | 69 | (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) | 
| 70 | end) cargs (bs, ts ~~ Ts)))) | |
| 71 | | _ => fold (add_binders thy i) ts bs) | |
| 72 | | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | |
| 73 | | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | |
| 74 | | add_binders thy i _ bs = bs; | |
| 75 | ||
| 69597 | 76 | fun split_conj f names (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ = (case head_of p of | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 77 | Const (name, _) => | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36428diff
changeset | 78 | if member (op =) names name then SOME (f p q) else NONE | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 79 | | _ => NONE) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 80 | | split_conj _ _ _ _ = NONE; | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 81 | |
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 82 | fun strip_all [] t = t | 
| 69597 | 83 | | strip_all (_ :: xs) (Const (\<^const_name>\<open>All\<close>, _) $ Abs (s, T, t)) = strip_all xs t; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 84 | |
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 85 | (*********************************************************************) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 86 | (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) | 
| 27722 
d657036e26c5
- corrected bogus comment for function inst_conj_all
 berghofe parents: 
27449diff
changeset | 87 | (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) | 
| 
d657036e26c5
- corrected bogus comment for function inst_conj_all
 berghofe parents: 
27449diff
changeset | 88 | (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) | 
| 
d657036e26c5
- corrected bogus comment for function inst_conj_all
 berghofe parents: 
27449diff
changeset | 89 | (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 90 | (* *) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 91 | (* where "id" protects the subformula from simplification *) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 92 | (*********************************************************************) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 93 | |
| 69597 | 94 | fun inst_conj_all names ps pis (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) _ = | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 95 | (case head_of p of | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 96 | Const (name, _) => | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36428diff
changeset | 97 | if member (op =) names name then SOME (HOLogic.mk_conj (p, | 
| 69597 | 98 | Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $ | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 99 | (subst_bounds (pis, strip_all pis q)))) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 100 | else NONE | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 101 | | _ => NONE) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 102 | | inst_conj_all names ps pis t u = | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 103 | if member (op aconv) ps (head_of u) then | 
| 69597 | 104 | SOME (Const (\<^const_name>\<open>Fun.id\<close>, HOLogic.boolT --> HOLogic.boolT) $ | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 105 | (subst_bounds (pis, strip_all pis t))) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 106 | else NONE | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 107 | | inst_conj_all _ _ _ _ _ = NONE; | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 108 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 109 | fun inst_conj_all_tac ctxt k = EVERY | 
| 60754 | 110 | [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), | 
| 111 | REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 112 |    simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1];
 | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 113 | |
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 114 | fun map_term f t u = (case f t u of | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 115 | NONE => map_term' f t u | x => x) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 116 | and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 117 | (NONE, NONE) => NONE | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 118 | | (SOME t'', NONE) => SOME (t'' $ u) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 119 | | (NONE, SOME u'') => SOME (t $ u'') | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 120 | | (SOME t'', SOME u'') => SOME (t'' $ u'')) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 121 | | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 122 | NONE => NONE | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 123 | | SOME t'' => SOME (Abs (s, T, t''))) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 124 | | map_term' _ _ _ = NONE; | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 125 | |
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 126 | (*********************************************************************) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 127 | (* Prove F[f t] from F[t], where F is monotone *) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 128 | (*********************************************************************) | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 129 | |
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 130 | fun map_thm ctxt f tac monos opt th = | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 131 | let | 
| 59582 | 132 | val prop = Thm.prop_of th; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 133 | fun prove t = | 
| 74519 | 134 |       Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} =>
 | 
| 135 | EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1, | |
| 136 | REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)), | |
| 137 | REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))]) | |
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 138 | in Option.map prove (map_term f prop (the_default prop opt)) end; | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 139 | |
| 27352 
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
 berghofe parents: 
27228diff
changeset | 140 | val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion; | 
| 
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
 berghofe parents: 
27228diff
changeset | 141 | |
| 25824 | 142 | fun first_order_matchs pats objs = Thm.first_order_match | 
| 27352 
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
 berghofe parents: 
27228diff
changeset | 143 | (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), | 
| 
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
 berghofe parents: 
27228diff
changeset | 144 | eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); | 
| 25824 | 145 | |
| 146 | fun first_order_mrs ths th = ths MRS | |
| 59582 | 147 | Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th; | 
| 25824 | 148 | |
| 74522 | 149 | fun prove_strong_ind s avoids lthy = | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 150 | let | 
| 74522 | 151 | val thy = Proof_Context.theory_of lthy; | 
| 25824 | 152 |     val ({names, ...}, {raw_induct, intrs, elims, ...}) =
 | 
| 74522 | 153 | Inductive.the_inductive_global lthy (Sign.intern_const thy s); | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 154 | val ind_params = Inductive.params_of raw_induct; | 
| 74522 | 155 | val raw_induct = atomize_induct lthy raw_induct; | 
| 156 | val elims = map (atomize_induct lthy) elims; | |
| 157 | val monos = Inductive.get_monos lthy; | |
| 158 | val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy; | |
| 33040 | 159 | val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of | 
| 22788 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 160 | [] => () | 
| 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 161 |       | xs => error ("Missing equivariance theorem for predicate(s): " ^
 | 
| 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 162 | commas_quote xs)); | 
| 44045 
2814ff2a6e3e
infrastructure for attaching names to hypothesis in cases; realised via the same tag mechanism as case names
 nipkow parents: 
43326diff
changeset | 163 | val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the | 
| 74522 | 164 | (Induct.lookup_inductP lthy (hd names))))); | 
| 165 | val (raw_induct', ctxt') = lthy | |
| 70326 | 166 | |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); | 
| 22530 | 167 | val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> | 
| 168 | HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); | |
| 169 | val ps = map (fst o snd) concls; | |
| 170 | ||
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58112diff
changeset | 171 | val _ = (case duplicates (op = o apply2 fst) avoids of | 
| 22530 | 172 | [] => () | 
| 173 |       | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
 | |
| 57884 
36b5691b81a5
tuned -- avoid confusion with @{assert} for system failures (exception Fail);
 wenzelm parents: 
56253diff
changeset | 174 | val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse | 
| 
36b5691b81a5
tuned -- avoid confusion with @{assert} for system failures (exception Fail);
 wenzelm parents: 
56253diff
changeset | 175 |       error ("Duplicate variable names for case " ^ quote a));
 | 
| 33040 | 176 | val _ = (case subtract (op =) induct_cases (map fst avoids) of | 
| 22530 | 177 | [] => () | 
| 178 |       | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
 | |
| 25824 | 179 |     val avoids' = if null induct_cases then replicate (length intrs) ("", [])
 | 
| 180 | else map (fn name => | |
| 181 | (name, the_default [] (AList.lookup op = avoids name))) induct_cases; | |
| 22530 | 182 | fun mk_avoids params (name, ps) = | 
| 183 | let val k = length params - 1 | |
| 184 | in map (fn x => case find_index (equal x o fst) params of | |
| 185 |           ~1 => error ("No such variable in case " ^ quote name ^
 | |
| 186 | " of inductive definition: " ^ quote x) | |
| 187 | | i => (Bound (k - i), snd (nth params i))) ps | |
| 188 | end; | |
| 189 | ||
| 190 | val prems = map (fn (prem, avoid) => | |
| 191 | let | |
| 192 | val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); | |
| 193 | val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); | |
| 194 | val params = Logic.strip_params prem | |
| 195 | in | |
| 196 | (params, | |
| 197 | fold (add_binders thy 0) (prems @ [concl]) [] @ | |
| 198 | map (apfst (incr_boundvars 1)) (mk_avoids params avoid), | |
| 199 | prems, strip_comb (HOLogic.dest_Trueprop concl)) | |
| 200 | end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); | |
| 201 | ||
| 202 | val atomTs = distinct op = (maps (map snd o #2) prems); | |
| 69597 | 203 | val ind_sort = if null atomTs then \<^sort>\<open>type\<close> | 
| 36428 
874843c1e96e
really minimize sorts after certification -- looks like this is intended here;
 wenzelm parents: 
36323diff
changeset | 204 | else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy | 
| 
874843c1e96e
really minimize sorts after certification -- looks like this is intended here;
 wenzelm parents: 
36323diff
changeset | 205 |         ("fs_" ^ Long_Name.base_name (fst (dest_Type T)))) atomTs));
 | 
| 43326 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 wenzelm parents: 
42361diff
changeset | 206 | val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 207 | val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; | 
| 22530 | 208 | val fsT = TFree (fs_ctxt_tyname, ind_sort); | 
| 209 | ||
| 60801 | 210 | val inductive_forall_def' = Thm.instantiate' | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59586diff
changeset | 211 | [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 212 | |
| 22530 | 213 | fun lift_pred' t (Free (s, T)) ts = | 
| 214 | list_comb (Free (s, fsT --> T), t :: ts); | |
| 215 | val lift_pred = lift_pred' (Bound 0); | |
| 216 | ||
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 217 | fun lift_prem (t as (f $ u)) = | 
| 22530 | 218 | let val (p, ts) = strip_comb t | 
| 219 | in | |
| 44868 | 220 | if member (op =) ps p then HOLogic.mk_induct_forall fsT $ | 
| 221 |               Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts))
 | |
| 22530 | 222 | else lift_prem f $ lift_prem u | 
| 223 | end | |
| 224 | | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | |
| 225 | | lift_prem t = t; | |
| 226 | ||
| 227 | fun mk_distinct [] = [] | |
| 32952 | 228 | | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) => | 
| 22530 | 229 | if T = U then SOME (HOLogic.mk_Trueprop | 
| 230 | (HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) | |
| 231 | else NONE) xs @ mk_distinct xs; | |
| 232 | ||
| 233 | fun mk_fresh (x, T) = HOLogic.mk_Trueprop | |
| 31938 | 234 | (NominalDatatype.fresh_const T fsT $ x $ Bound 0); | 
| 22530 | 235 | |
| 236 | val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => | |
| 237 | let | |
| 238 |         val params' = params @ [("y", fsT)];
 | |
| 239 | val prem = Logic.list_implies | |
| 240 | (map mk_fresh bvars @ mk_distinct bvars @ | |
| 241 | map (fn prem => | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 242 | if null (preds_of ps prem) then prem | 
| 22530 | 243 | else lift_prem prem) prems, | 
| 244 | HOLogic.mk_Trueprop (lift_pred p ts)); | |
| 29276 | 245 | val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') | 
| 22530 | 246 | in | 
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
44929diff
changeset | 247 | (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem))) | 
| 22530 | 248 | end) prems); | 
| 249 | ||
| 250 | val ind_vars = | |
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
57884diff
changeset | 251 | (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ | 
| 22530 | 252 |        map NominalAtoms.mk_permT atomTs) @ [("z", fsT)];
 | 
| 253 | val ind_Ts = rev (map snd ind_vars); | |
| 254 | ||
| 255 | val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj | |
| 256 | (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, | |
| 257 | HOLogic.list_all (ind_vars, lift_pred p | |
| 31938 | 258 | (map (fold_rev (NominalDatatype.mk_perm ind_Ts) | 
| 22530 | 259 | (map Bound (length atomTs downto 1))) ts)))) concls)); | 
| 260 | ||
| 261 | val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj | |
| 262 | (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, | |
| 263 | lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); | |
| 264 | ||
| 265 | val vc_compat = map (fn (params, bvars, prems, (p, ts)) => | |
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
44929diff
changeset | 266 | map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies | 
| 32952 | 267 | (map_filter (fn prem => | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 268 | if null (preds_of ps prem) then SOME prem | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 269 | else map_term (split_conj (K o I) names) prem prem) prems, q)))) | 
| 22530 | 270 | (mk_distinct bvars @ | 
| 271 | maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop | |
| 31938 | 272 | (NominalDatatype.fresh_const U T $ u $ t)) bvars) | 
| 22530 | 273 | (ts ~~ binder_types (fastype_of p)))) prems; | 
| 274 | ||
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 275 | val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 276 | val pt2_atoms = map (fn aT => Global_Theory.get_thm thy | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 277 |       ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs;
 | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 278 | fun eqvt_ss ctxt = | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 279 | put_simpset HOL_basic_ss ctxt | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 280 | addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 281 | addsimprocs [mk_perm_bool_simproc [\<^const_name>\<open>Fun.id\<close>], | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 282 | NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 283 | val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 284 | val perm_bij = Global_Theory.get_thms thy "perm_bij"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 285 | val fs_atoms = map (fn aT => Global_Theory.get_thm thy | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 286 |       ("fs_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "1")) atomTs;
 | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 287 | val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 288 | val fresh_atm = Global_Theory.get_thms thy "fresh_atm"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 289 | val swap_simps = Global_Theory.get_thms thy "swap_simps"; | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 290 | val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh"; | 
| 22530 | 291 | |
| 292 | fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = | |
| 293 | let | |
| 27722 
d657036e26c5
- corrected bogus comment for function inst_conj_all
 berghofe parents: 
27449diff
changeset | 294 | (** protect terms to avoid that fresh_prod interferes with **) | 
| 22530 | 295 | (** pairs used in introduction rules of inductive predicate **) | 
| 296 | fun protect t = | |
| 69597 | 297 | let val T = fastype_of t in Const (\<^const_name>\<open>Fun.id\<close>, T --> T) $ t end; | 
| 22530 | 298 | val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); | 
| 299 | val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop | |
| 300 |             (HOLogic.exists_const T $ Abs ("x", T,
 | |
| 31938 | 301 | NominalDatatype.fresh_const T (fastype_of p) $ | 
| 22530 | 302 | Bound 0 $ p))) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 303 |           (fn {context = goal_ctxt, ...} => EVERY
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 304 | [resolve_tac goal_ctxt exists_fresh' 1, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 305 | resolve_tac goal_ctxt fs_atoms 1]); | 
| 32202 | 306 | val (([(_, cx)], ths), ctxt') = Obtain.result | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 307 | (fn goal_ctxt => EVERY | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 308 | [eresolve_tac goal_ctxt [exE] 1, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 309 | full_simp_tac (put_simpset HOL_ss goal_ctxt addsimps (fresh_prod :: fresh_atm)) 1, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 310 |              full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps [@{thm id_apply}]) 1,
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 311 | REPEAT (eresolve_tac goal_ctxt [conjE] 1)]) | 
| 22530 | 312 | [ex] ctxt | 
| 59582 | 313 | in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; | 
| 22530 | 314 | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 315 | fun mk_ind_proof ctxt thss = | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 316 |       Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 317 |         let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 318 | resolve_tac goal_ctxt1 [raw_induct] 1 THEN | 
| 22530 | 319 | EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 320 | [REPEAT (resolve_tac goal_ctxt1 [allI] 1), | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 321 | simp_tac (eqvt_ss goal_ctxt1) 1, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 322 |              SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} =>
 | 
| 22530 | 323 | let | 
| 324 | val (params', (pis, z)) = | |
| 59582 | 325 | chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||> | 
| 22530 | 326 | split_last; | 
| 327 | val bvars' = map | |
| 328 | (fn (Bound i, T) => (nth params' (length params' - i), T) | |
| 329 | | (t, T) => (t, T)) bvars; | |
| 330 | val pi_bvars = map (fn (t, _) => | |
| 31938 | 331 | fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; | 
| 59582 | 332 | val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 333 | val (freshs1, freshs2, ctxt') = fold | 
| 22530 | 334 | (obtain_fresh_name (ts @ pi_bvars)) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 335 | (map snd bvars') ([], [], goal_ctxt2); | 
| 31938 | 336 | val freshs2' = NominalDatatype.mk_not_sym freshs2; | 
| 337 | val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); | |
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 338 | fun concat_perm pi1 pi2 = | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 339 | let val T = fastype_of pi1 | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 340 | in if T = fastype_of pi2 then | 
| 69597 | 341 | Const (\<^const_name>\<open>append\<close>, T --> T --> T) $ pi1 $ pi2 | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 342 | else pi2 | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 343 | end; | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 344 | val pis'' = fold (concat_perm #> map) pis' pis; | 
| 59582 | 345 | val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) | 
| 22530 | 346 | (Vartab.empty, Vartab.empty); | 
| 74282 | 347 | val ihyp' = Thm.instantiate (TVars.empty, | 
| 348 | Vars.make (map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t)) | |
| 32035 | 349 | (map (Envir.subst_term env) vs ~~ | 
| 31938 | 350 | map (fold_rev (NominalDatatype.mk_perm []) | 
| 74282 | 351 | (rev pis' @ pis)) params' @ [z]))) ihyp; | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 352 | fun mk_pi th = | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 353 |                    Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]
 | 
| 31938 | 354 | addsimprocs [NominalDatatype.perm_simproc]) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 355 | (Simplifier.simplify (eqvt_ss ctxt') | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 356 | (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt') | 
| 25824 | 357 | (rev pis' @ pis) th)); | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 358 | val (gprems1, gprems2) = split_list | 
| 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 359 | (map (fn (th, t) => | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 360 | if null (preds_of ps t) then (SOME th, mk_pi th) | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 361 | else | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 362 | (map_thm ctxt' (split_conj (K o I) names) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 363 | (eresolve_tac ctxt' [conjunct1] 1) monos NONE th, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 364 | mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 365 | (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th)))) | 
| 32952 | 366 | (gprems ~~ oprems)) |>> map_filter I; | 
| 22530 | 367 | val vc_compat_ths' = map (fn th => | 
| 368 | let | |
| 25824 | 369 | val th' = first_order_mrs gprems1 th; | 
| 59582 | 370 | val (bop, lhs, rhs) = (case Thm.concl_of th' of | 
| 22530 | 371 | _ $ (fresh $ lhs $ rhs) => | 
| 372 | (fn t => fn u => fresh $ t $ u, lhs, rhs) | |
| 373 | | _ $ (_ $ (_ $ lhs $ rhs)) => | |
| 374 | (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 375 | val th'' = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop | 
| 31938 | 376 | (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) | 
| 377 | (fold_rev (NominalDatatype.mk_perm []) pis rhs))) | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 378 |                        (fn {context = goal_ctxt3, ...} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 379 | simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 380 | (fresh_bij @ perm_bij)) 1 THEN resolve_tac goal_ctxt3 [th'] 1) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 381 | in Simplifier.simplify (eqvt_ss ctxt' addsimps fresh_atm) th'' end) | 
| 22530 | 382 | vc_compat_ths; | 
| 31938 | 383 | val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; | 
| 27449 
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
 berghofe parents: 
27353diff
changeset | 384 | (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) | 
| 
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
 berghofe parents: 
27353diff
changeset | 385 | (** we have to pre-simplify the rewrite rules **) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 386 | val swap_simps_simpset = put_simpset HOL_ss ctxt' addsimps swap_simps @ | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 387 | map (Simplifier.simplify (put_simpset HOL_ss ctxt' addsimps swap_simps)) | 
| 22530 | 388 | (vc_compat_ths'' @ freshs2'); | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 389 | val th = Goal.prove ctxt' [] [] | 
| 22530 | 390 | (HOLogic.mk_Trueprop (list_comb (P $ hd ts, | 
| 31938 | 391 | map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 392 |                    (fn {context = goal_ctxt4, ...} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 393 | EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1, | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 394 | resolve_tac goal_ctxt4 [ihyp'] 1, | 
| 59582 | 395 | REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 396 | (simp_tac swap_simps_simpset 1), | 
| 22530 | 397 | REPEAT_DETERM_N (length gprems) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 398 | (simp_tac (put_simpset HOL_basic_ss goal_ctxt4 | 
| 27847 
0dffedf9aff5
Changed proof of strong induction rule to avoid infinite loop
 berghofe parents: 
27722diff
changeset | 399 | addsimps [inductive_forall_def'] | 
| 31938 | 400 | addsimprocs [NominalDatatype.perm_simproc]) 1 THEN | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 401 | resolve_tac goal_ctxt4 gprems2 1)])); | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 402 | val final = Goal.prove ctxt' [] [] (Thm.term_of concl) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 403 |                    (fn {context = goal_ctxt5, ...} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 404 | cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5 | 
| 22530 | 405 | addsimps vc_compat_ths'' @ freshs2' @ | 
| 406 | perm_fresh_fresh @ fresh_atm) 1); | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 407 | val final' = Proof_Context.export ctxt' goal_ctxt2 [final]; | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 408 | in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1]) | 
| 22530 | 409 | (prems ~~ thss ~~ ihyps ~~ prems''))) | 
| 410 | in | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 411 | cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 412 | REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 413 | eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 414 |             REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 415 | asm_full_simp_tac goal_ctxt 1) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 416 | end) |> singleton (Proof_Context.export ctxt lthy); | 
| 22530 | 417 | |
| 25824 | 418 | (** strong case analysis rule **) | 
| 419 | ||
| 420 | val cases_prems = map (fn ((name, avoids), rule) => | |
| 421 | let | |
| 74522 | 422 | val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] lthy; | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 423 | val prem :: prems = Logic.strip_imp_prems rule'; | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 424 | val concl = Logic.strip_imp_concl rule' | 
| 25824 | 425 | in | 
| 426 | (prem, | |
| 427 | List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), | |
| 428 | concl, | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 429 | fold_map (fn (prem, (_, avoid)) => fn ctxt => | 
| 25824 | 430 | let | 
| 431 | val prems = Logic.strip_assums_hyp prem; | |
| 432 | val params = Logic.strip_params prem; | |
| 433 | val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 434 | fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) = | 
| 25824 | 435 | if member (op = o apsnd fst) bnds (Bound i) then | 
| 436 | let | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 437 | val ([s'], ctxt') = Variable.variant_fixes [s] ctxt; | 
| 25824 | 438 | val t = Free (s', T) | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 439 | in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 440 | else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts); | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 441 | val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 442 | (0, 0, ctxt, [], [], [], []) | 
| 25824 | 443 | in | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 444 | ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt') | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 445 | end) (prems ~~ avoids) ctxt') | 
| 25824 | 446 | end) | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 447 | (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~ | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 448 | elims); | 
| 25824 | 449 | |
| 450 | val cases_prems' = | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 451 | map (fn (prem, args, concl, (prems, _)) => | 
| 25824 | 452 | let | 
| 453 | fun mk_prem (ps, [], _, prems) = | |
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
44929diff
changeset | 454 | Logic.list_all (ps, Logic.list_implies (prems, concl)) | 
| 25824 | 455 | | mk_prem (ps, qs, _, prems) = | 
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
44929diff
changeset | 456 | Logic.list_all (ps, Logic.mk_implies | 
| 25824 | 457 | (Logic.list_implies | 
| 458 | (mk_distinct qs @ | |
| 459 | maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop | |
| 31938 | 460 | (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) | 
| 25824 | 461 | args) qs, | 
| 462 | HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj | |
| 463 | (map HOLogic.dest_Trueprop prems))), | |
| 464 | concl)) | |
| 465 | in map mk_prem prems end) cases_prems; | |
| 466 | ||
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 467 | fun cases_eqvt_simpset ctxt = | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 468 | put_simpset HOL_ss ctxt | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 469 | addsimps eqvt_thms @ swap_simps @ perm_pi_simp | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 470 | addsimprocs [NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; | 
| 27352 
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
 berghofe parents: 
27228diff
changeset | 471 | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 472 | fun simp_fresh_atm ctxt = | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 473 | Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps fresh_atm); | 
| 25824 | 474 | |
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 475 | fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt))), | 
| 25824 | 476 | prems') = | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 477 | (name, Goal.prove ctxt [] (prem :: prems') concl | 
| 26711 | 478 |         (fn {prems = hyp :: hyps, context = ctxt1} =>
 | 
| 60754 | 479 | EVERY (resolve_tac ctxt1 [hyp RS elim] 1 :: | 
| 25824 | 480 | map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => | 
| 481 |             SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} =>
 | |
| 482 | if null qs then | |
| 60754 | 483 | resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1 | 
| 25824 | 484 | else | 
| 485 | let | |
| 59582 | 486 | val params' = map (Thm.term_of o #2 o nth (rev params)) is; | 
| 25824 | 487 | val tab = params' ~~ map fst qs; | 
| 488 | val (hyps1, hyps2) = chop (length args) case_hyps; | |
| 489 | (* turns a = t and [x1 # t, ..., xn # t] *) | |
| 490 | (* into [x1 # a, ..., xn # a] *) | |
| 491 | fun inst_fresh th' ths = | |
| 492 | let val (ths1, ths2) = chop (length qs) ths | |
| 493 | in | |
| 494 | (map (fn th => | |
| 495 | let | |
| 496 | val (cf, ct) = | |
| 59582 | 497 | Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)); | 
| 60801 | 498 | val arg_cong' = Thm.instantiate' | 
| 59586 | 499 | [SOME (Thm.ctyp_of_cterm ct)] | 
| 25824 | 500 | [NONE, SOME ct, SOME cf] (arg_cong RS iffD2); | 
| 501 | val inst = Thm.first_order_match (ct, | |
| 59582 | 502 | Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th'))) | 
| 25824 | 503 | in [th', th] MRS Thm.instantiate inst arg_cong' | 
| 504 | end) ths1, | |
| 505 | ths2) | |
| 506 | end; | |
| 507 | val (vc_compat_ths1, vc_compat_ths2) = | |
| 508 | chop (length vc_compat_ths - length args * length qs) | |
| 509 | (map (first_order_mrs hyps2) vc_compat_ths); | |
| 510 | val vc_compat_ths' = | |
| 31938 | 511 | NominalDatatype.mk_not_sym vc_compat_ths1 @ | 
| 25824 | 512 | flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); | 
| 513 | val (freshs1, freshs2, ctxt3) = fold | |
| 514 | (obtain_fresh_name (args @ map fst qs @ params')) | |
| 515 | (map snd qs) ([], [], ctxt2); | |
| 31938 | 516 | val freshs2' = NominalDatatype.mk_not_sym freshs2; | 
| 517 | val pis = map (NominalDatatype.perm_of_pair) | |
| 25824 | 518 | ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); | 
| 60787 | 519 | val mk_pis = fold_rev (mk_perm_bool ctxt3) (map (Thm.cterm_of ctxt3) pis); | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59586diff
changeset | 520 | val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms | 
| 25824 | 521 | (fn x as Free _ => | 
| 36692 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 haftmann parents: 
36428diff
changeset | 522 | if member (op =) args x then x | 
| 25824 | 523 | else (case AList.lookup op = tab x of | 
| 524 | SOME y => y | |
| 31938 | 525 | | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) | 
| 59582 | 526 | | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps)); | 
| 25824 | 527 | val inst = Thm.first_order_match (Thm.dest_arg | 
| 528 | (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); | |
| 59582 | 529 | val th = Goal.prove ctxt3 [] [] (Thm.term_of concl) | 
| 25824 | 530 |                     (fn {context = ctxt4, ...} =>
 | 
| 60754 | 531 | resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN | 
| 532 |                        SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} =>
 | |
| 25824 | 533 | let | 
| 31938 | 534 | val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 535 | val case_simpset = cases_eqvt_simpset ctxt5 addsimps freshs2' @ | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 536 | map (simp_fresh_atm ctxt5) (vc_compat_ths' @ fresh_hyps'); | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 537 | val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; | 
| 25824 | 538 | val hyps1' = map | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 539 | (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; | 
| 25824 | 540 | val hyps2' = map | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 541 | (mk_pis #> Simplifier.simplify case_simpset) hyps2; | 
| 27449 
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
 berghofe parents: 
27353diff
changeset | 542 | val case_hyps' = hyps1' @ hyps2' | 
| 25824 | 543 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50771diff
changeset | 544 | simp_tac case_simpset 1 THEN | 
| 60754 | 545 | REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN | 
| 546 | resolve_tac ctxt5 case_hyps' 1) | |
| 25824 | 547 | end) ctxt4 1) | 
| 42361 | 548 | val final = Proof_Context.export ctxt3 ctxt2 [th] | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 549 | in resolve_tac ctxt2 final 1 end) ctxt1 1) | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 550 | (thss ~~ hyps ~~ prems))) |> | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 551 | singleton (Proof_Context.export ctxt lthy)) | 
| 25824 | 552 | |
| 22530 | 553 | in | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 554 | ctxt'' |> | 
| 74522 | 555 | Proof.theorem NONE (fn thss => fn lthy1 => | 
| 22530 | 556 | let | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 557 | val rec_name = space_implode "_" (map Long_Name.base_name names); | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30087diff
changeset | 558 | val rec_qualified = Binding.qualify false rec_name; | 
| 33368 | 559 | val ind_case_names = Rule_Cases.case_names induct_cases; | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 560 | val induct_cases' = Inductive.partition_rules' raw_induct | 
| 25824 | 561 | (intrs ~~ induct_cases); | 
| 74522 | 562 | val thss' = map (map (atomize_intr lthy1)) thss; | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 563 | val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 564 | val strong_raw_induct = | 
| 74522 | 565 | mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1; | 
| 566 | val strong_cases = map (mk_cases_proof ##> Inductive.rulify lthy1) | |
| 25824 | 567 | (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); | 
| 50771 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 568 | val strong_induct_atts = | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 569 | map (Attrib.internal o K) | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 570 | [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; | 
| 22530 | 571 | val strong_induct = | 
| 50771 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 572 | if length names > 1 then strong_raw_induct | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 573 | else strong_raw_induct RSN (2, rev_mp); | 
| 74522 | 574 | val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note | 
| 50771 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 575 | ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]); | 
| 22530 | 576 | val strong_inducts = | 
| 74522 | 577 | Project_Rule.projects lthy1 (1 upto length names) strong_induct'; | 
| 22530 | 578 | in | 
| 74522 | 579 | lthy2 |> | 
| 50771 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 580 | Local_Theory.notes | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 581 | [((rec_qualified (Binding.name "strong_inducts"), []), | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 582 | strong_inducts |> map (fn th => ([th], | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 583 | [Attrib.internal (K ind_case_names), | 
| 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 584 | Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> | 
| 33671 | 585 | Local_Theory.notes (map (fn ((name, elim), (_, cases)) => | 
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 586 | ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), | 
| 33368 | 587 | [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), | 
| 50771 
2852f997bfb5
prefer negative "consumes", relative to the total number of prems, which is stable under more morphisms, notably those from nested context with assumes (cf. existing treatment of 'obtains');
 wenzelm parents: 
46961diff
changeset | 588 | Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 589 | (strong_cases ~~ induct_cases')) |> snd | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 590 | end) | 
| 24570 
621b60b1df00
Generalized equivariance and nominal_inductive commands to
 berghofe parents: 
23531diff
changeset | 591 | (map (map (rulify_term thy #> rpair [])) vc_compat) | 
| 22530 | 592 | end; | 
| 593 | ||
| 74519 | 594 | fun prove_eqvt s xatoms lthy = | 
| 22530 | 595 | let | 
| 74519 | 596 | val thy = Proof_Context.theory_of lthy; | 
| 22788 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 597 |     val ({names, ...}, {raw_induct, intrs, elims, ...}) =
 | 
| 74519 | 598 | Inductive.the_inductive_global lthy (Sign.intern_const thy s); | 
| 599 | val raw_induct = atomize_induct lthy raw_induct; | |
| 600 | val elims = map (atomize_induct lthy) elims; | |
| 601 | val intrs = map (atomize_intr lthy) intrs; | |
| 602 | val monos = Inductive.get_monos lthy; | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 603 | val intrs' = Inductive.unpartition_rules intrs | 
| 22788 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 604 | (map (fn (((s, ths), (_, k)), th) => | 
| 60362 | 605 | (s, ths ~~ Inductive.infer_intro_vars thy th k ths)) | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 606 | (Inductive.partition_rules raw_induct intrs ~~ | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 607 | Inductive.arities_of raw_induct ~~ elims)); | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 608 | val k = length (Inductive.params_of raw_induct); | 
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 609 | val atoms' = NominalAtoms.atoms_of thy; | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 610 | val atoms = | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 611 | if null xatoms then atoms' else | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 612 | let val atoms = map (Sign.intern_type thy) xatoms | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 613 | in | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 614 | (case duplicates op = atoms of | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 615 | [] => () | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 616 |            | xs => error ("Duplicate atoms: " ^ commas xs);
 | 
| 33040 | 617 | case subtract (op =) atoms' atoms of | 
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 618 | [] => () | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 619 |            | xs => error ("No such atoms: " ^ commas xs);
 | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 620 | atoms) | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22544diff
changeset | 621 | end; | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39159diff
changeset | 622 | val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; | 
| 74519 | 623 | val (([t], [pi]), ctxt1) = lthy |> | 
| 59582 | 624 | Variable.import_terms false [Thm.concl_of raw_induct] ||>> | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 625 | Variable.variant_fixes ["pi"]; | 
| 74519 | 626 | fun eqvt_simpset ctxt = put_simpset HOL_basic_ss ctxt addsimps | 
| 627 | (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs | |
| 54991 | 628 | [mk_perm_bool_simproc names, | 
| 629 | NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; | |
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 630 | val ps = map (fst o HOLogic.dest_imp) | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 631 | (HOLogic.dest_conj (HOLogic.dest_Trueprop t)); | 
| 74519 | 632 | fun eqvt_tac ctxt pi (intr, vs) st = | 
| 22544 | 633 | let | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 634 | fun eqvt_err s = | 
| 74519 | 635 | let val ([t], ctxt') = Variable.import_terms true [Thm.prop_of intr] ctxt | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 636 |           in error ("Could not prove equivariance for introduction rule\n" ^
 | 
| 74519 | 637 | Syntax.string_of_term ctxt' t ^ "\n" ^ s) | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 638 | end; | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 639 |         val res = SUBPROOF (fn {context = goal_ctxt, prems, params, ...} =>
 | 
| 22788 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 640 | let | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 641 | val prems' = map (fn th => the_default th (map_thm goal_ctxt | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 642 | (split_conj (K I) names) (eresolve_tac goal_ctxt [conjunct2] 1) monos NONE th)) prems; | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 643 | val prems'' = map (fn th => Simplifier.simplify (eqvt_simpset goal_ctxt) | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 644 | (mk_perm_bool goal_ctxt (Thm.cterm_of goal_ctxt pi) th)) prems'; | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 645 | val intr' = infer_instantiate goal_ctxt (map (#1 o dest_Var) vs ~~ | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 646 | map (Thm.cterm_of goal_ctxt o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params) | 
| 22788 
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
 berghofe parents: 
22755diff
changeset | 647 | intr | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 648 | in (resolve_tac goal_ctxt [intr'] THEN_ALL_NEW (TRY o resolve_tac goal_ctxt prems'')) 1 | 
| 74519 | 649 | end) ctxt 1 st | 
| 22544 | 650 | in | 
| 651 | case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of | |
| 652 |           NONE => eqvt_err ("Rule does not match goal\n" ^
 | |
| 74519 | 653 | Syntax.string_of_term ctxt (hd (Thm.prems_of st))) | 
| 22544 | 654 | | SOME (th, _) => Seq.single th | 
| 655 | end; | |
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 656 | val thss = map (fn atom => | 
| 25824 | 657 | let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) | 
| 22530 | 658 | in map (fn th => zero_var_indexes (th RS mp)) | 
| 74519 | 659 | (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt1 [] [] | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 660 | (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 661 | let | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 662 | val (h, ts) = strip_comb p; | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 663 | val (ts1, ts2) = chop k ts | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 664 | in | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 665 | HOLogic.mk_imp (p, list_comb (h, ts1 @ | 
| 31938 | 666 | map (NominalDatatype.mk_perm [] pi') ts2)) | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 667 | end) ps))) | 
| 74524 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 668 |           (fn {context = goal_ctxt, ...} =>
 | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 669 | EVERY (resolve_tac goal_ctxt [raw_induct] 1 :: map (fn intr_vs => | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 670 | full_simp_tac (eqvt_simpset goal_ctxt) 1 THEN | 
| 
8ee3d5d3c1bf
more accurate treatment of context, notably for nested Goal.proof / SUBPROOF;
 wenzelm parents: 
74522diff
changeset | 671 | eqvt_tac goal_ctxt pi' intr_vs) intrs')) |> | 
| 74519 | 672 | singleton (Proof_Context.export ctxt1 lthy))) | 
| 22544 | 673 | end) atoms | 
| 674 | in | |
| 74519 | 675 | lthy |> | 
| 33671 | 676 | Local_Theory.notes (map (fn (name, ths) => | 
| 30435 
e62d6ecab6ad
explicit Binding.qualified_name -- prevents implicitly qualified bstring;
 wenzelm parents: 
30364diff
changeset | 677 | ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 678 | [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) | 
| 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 679 | (names ~~ transp thss)) |> snd | 
| 22544 | 680 | end; | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 681 | |
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 682 | |
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 683 | (* outer syntax *) | 
| 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 684 | |
| 24867 | 685 | val _ = | 
| 69597 | 686 | Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>nominal_inductive\<close> | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36692diff
changeset | 687 | "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" | 
| 69597 | 688 | (Parse.name -- Scan.optional (\<^keyword>\<open>avoids\<close> |-- Parse.and_list1 (Parse.name -- | 
| 689 | (\<^keyword>\<open>:\<close> |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) => | |
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 690 | prove_strong_ind name avoids)); | 
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 691 | |
| 24867 | 692 | val _ = | 
| 69597 | 693 | Outer_Syntax.local_theory \<^command_keyword>\<open>equivariance\<close> | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 694 | "prove equivariance for inductive predicate involving nominal datatypes" | 
| 69597 | 695 | (Parse.name -- Scan.optional (\<^keyword>\<open>[\<close> |-- Parse.list1 Parse.name --| \<^keyword>\<open>]\<close>) [] >> | 
| 30087 
a780642a9c9c
nominal_inductive and equivariance now work on local_theory.
 berghofe parents: 
29585diff
changeset | 696 | (fn (name, atoms) => prove_eqvt name atoms)); | 
| 22530 | 697 | |
| 22313 
1a507b463f50
First steps towards strengthening of induction rules for
 berghofe parents: diff
changeset | 698 | end |