author | wenzelm |
Wed, 31 Dec 2008 19:54:03 +0100 | |
changeset 29276 | 94b1ffec9201 |
parent 29270 | 0eade173f77e |
child 29281 | b22ccb3998db |
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 |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
10 |
val prove_strong_ind: string -> (string * string list) list -> theory -> Proof.state |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
11 |
val prove_eqvt: string -> string list -> theory -> 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 |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
17 |
val inductive_forall_name = "HOL.induct_forall"; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
18 |
val inductive_forall_def = thm "induct_forall_def"; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
19 |
val inductive_atomize = thms "induct_atomize"; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
20 |
val inductive_rulify = thms "induct_rulify"; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
21 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
22 |
fun rulify_term thy = MetaSimplifier.rewrite_term thy inductive_rulify []; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
23 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
24 |
val atomize_conv = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
25 |
MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE)) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
26 |
(HOL_basic_ss addsimps inductive_atomize); |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
27 |
val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv); |
24832 | 28 |
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 |
26568 | 29 |
(Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt)); |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
30 |
|
22530 | 31 |
val fresh_prod = thm "fresh_prod"; |
32 |
||
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
33 |
val perm_bool = mk_meta_eq (thm "perm_bool"); |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
34 |
val perm_boolI = thm "perm_boolI"; |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
35 |
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
36 |
(Drule.strip_imp_concl (cprop_of perm_boolI)))); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
37 |
|
25824 | 38 |
fun mk_perm_bool pi th = th RS Drule.cterm_instantiate |
39 |
[(perm_boolI_pi, pi)] perm_boolI; |
|
40 |
||
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
41 |
fun mk_perm_bool_simproc names = Simplifier.simproc_i |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
42 |
(theory_of_thm perm_bool) "perm_bool" [@{term "perm pi x"}] (fn thy => fn ss => |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
43 |
fn Const ("Nominal.perm", _) $ _ $ t => |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
44 |
if the_default "" (try (head_of #> dest_Const #> fst) t) mem names |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
45 |
then SOME perm_bool else NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
46 |
| _ => NONE); |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
47 |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
48 |
fun transp ([] :: _) = [] |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
49 |
| transp xs = map hd xs :: transp (map tl xs); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
50 |
|
22530 | 51 |
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of |
52 |
(Const (s, T), ts) => (case strip_type T of |
|
53 |
(Ts, Type (tname, _)) => |
|
54 |
(case NominalPackage.get_nominal_datatype thy tname of |
|
55 |
NONE => fold (add_binders thy i) ts bs |
|
56 |
| SOME {descr, index, ...} => (case AList.lookup op = |
|
57 |
(#3 (the (AList.lookup op = descr index))) s of |
|
58 |
NONE => fold (add_binders thy i) ts bs |
|
59 |
| SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => |
|
60 |
let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' |
|
61 |
in (add_binders thy i u |
|
62 |
(fold (fn (u, T) => |
|
63 |
if exists (fn j => j < i) (loose_bnos u) then I |
|
64 |
else insert (op aconv o pairself fst) |
|
65 |
(incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) |
|
66 |
end) cargs (bs, ts ~~ Ts)))) |
|
67 |
| _ => fold (add_binders thy i) ts bs) |
|
68 |
| (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) |
|
69 |
| add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs |
|
70 |
| add_binders thy i _ bs = bs; |
|
71 |
||
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
72 |
fun split_conj f names (Const ("op &", _) $ p $ q) _ = (case head_of p of |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
73 |
Const (name, _) => |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
74 |
if name mem names then SOME (f p q) else NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
75 |
| _ => NONE) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
76 |
| split_conj _ _ _ _ = NONE; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
77 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
78 |
fun strip_all [] t = t |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
79 |
| strip_all (_ :: xs) (Const ("All", _) $ Abs (s, T, t)) = strip_all xs t; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
80 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
81 |
(*********************************************************************) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
82 |
(* 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:
27449
diff
changeset
|
83 |
(* 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:
27449
diff
changeset
|
84 |
(* 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:
27449
diff
changeset
|
85 |
(* 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:
23531
diff
changeset
|
86 |
(* *) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
87 |
(* where "id" protects the subformula from simplification *) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
88 |
(*********************************************************************) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
89 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
90 |
fun inst_conj_all names ps pis (Const ("op &", _) $ p $ q) _ = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
91 |
(case head_of p of |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
92 |
Const (name, _) => |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
93 |
if name mem names then SOME (HOLogic.mk_conj (p, |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
94 |
Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $ |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
95 |
(subst_bounds (pis, strip_all pis q)))) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
96 |
else NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
97 |
| _ => NONE) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
98 |
| inst_conj_all names ps pis t u = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
99 |
if member (op aconv) ps (head_of u) then |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
100 |
SOME (Const ("Fun.id", HOLogic.boolT --> HOLogic.boolT) $ |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
101 |
(subst_bounds (pis, strip_all pis t))) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
102 |
else NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
103 |
| inst_conj_all _ _ _ _ _ = NONE; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
104 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
105 |
fun inst_conj_all_tac k = EVERY |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
106 |
[TRY (EVERY [etac conjE 1, rtac conjI 1, atac 1]), |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
107 |
REPEAT_DETERM_N k (etac allE 1), |
26359 | 108 |
simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1]; |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
109 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
110 |
fun map_term f t u = (case f t u of |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
111 |
NONE => map_term' f t u | x => x) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
112 |
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:
23531
diff
changeset
|
113 |
(NONE, NONE) => NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
114 |
| (SOME t'', NONE) => SOME (t'' $ u) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
115 |
| (NONE, SOME u'') => SOME (t $ u'') |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
116 |
| (SOME t'', SOME u'') => SOME (t'' $ u'')) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
117 |
| 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:
23531
diff
changeset
|
118 |
NONE => NONE |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
119 |
| SOME t'' => SOME (Abs (s, T, t''))) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
120 |
| map_term' _ _ _ = NONE; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
121 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
122 |
(*********************************************************************) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
123 |
(* Prove F[f t] from F[t], where F is monotone *) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
124 |
(*********************************************************************) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
125 |
|
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
126 |
fun map_thm ctxt f tac monos opt th = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
127 |
let |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
128 |
val prop = prop_of th; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
129 |
fun prove t = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
130 |
Goal.prove ctxt [] [] t (fn _ => |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
131 |
EVERY [cut_facts_tac [th] 1, etac rev_mp 1, |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
132 |
REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)), |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
133 |
REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
134 |
in Option.map prove (map_term f prop (the_default prop opt)) end; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
135 |
|
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
136 |
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:
27228
diff
changeset
|
137 |
|
25824 | 138 |
fun first_order_matchs pats objs = Thm.first_order_match |
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
139 |
(eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
140 |
eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); |
25824 | 141 |
|
142 |
fun first_order_mrs ths th = ths MRS |
|
143 |
Thm.instantiate (first_order_matchs (cprems_of th) (map cprop_of ths)) th; |
|
144 |
||
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
145 |
fun prove_strong_ind s avoids thy = |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
146 |
let |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
147 |
val ctxt = ProofContext.init thy; |
25824 | 148 |
val ({names, ...}, {raw_induct, intrs, elims, ...}) = |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
149 |
InductivePackage.the_inductive ctxt (Sign.intern_const thy s); |
25824 | 150 |
val ind_params = InductivePackage.params_of raw_induct; |
24832 | 151 |
val raw_induct = atomize_induct ctxt raw_induct; |
25824 | 152 |
val elims = map (atomize_induct ctxt) elims; |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
153 |
val monos = InductivePackage.get_monos ctxt; |
24571 | 154 |
val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt; |
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
155 |
val _ = (case names \\ foldl (apfst prop_of #> OldTerm.add_term_consts) [] eqvt_thms of |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
156 |
[] => () |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
157 |
| xs => error ("Missing equivariance theorem for predicate(s): " ^ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
158 |
commas_quote xs)); |
22530 | 159 |
val induct_cases = map fst (fst (RuleCases.get (the |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24832
diff
changeset
|
160 |
(Induct.lookup_inductP ctxt (hd names))))); |
22530 | 161 |
val raw_induct' = Logic.unvarify (prop_of raw_induct); |
25824 | 162 |
val elims' = map (Logic.unvarify o prop_of) elims; |
22530 | 163 |
val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> |
164 |
HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); |
|
165 |
val ps = map (fst o snd) concls; |
|
166 |
||
167 |
val _ = (case duplicates (op = o pairself fst) avoids of |
|
168 |
[] => () |
|
169 |
| xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); |
|
170 |
val _ = assert_all (null o duplicates op = o snd) avoids |
|
171 |
(fn (a, _) => error ("Duplicate variable names for case " ^ quote a)); |
|
172 |
val _ = (case map fst avoids \\ induct_cases of |
|
173 |
[] => () |
|
174 |
| xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); |
|
25824 | 175 |
val avoids' = if null induct_cases then replicate (length intrs) ("", []) |
176 |
else map (fn name => |
|
177 |
(name, the_default [] (AList.lookup op = avoids name))) induct_cases; |
|
22530 | 178 |
fun mk_avoids params (name, ps) = |
179 |
let val k = length params - 1 |
|
180 |
in map (fn x => case find_index (equal x o fst) params of |
|
181 |
~1 => error ("No such variable in case " ^ quote name ^ |
|
182 |
" of inductive definition: " ^ quote x) |
|
183 |
| i => (Bound (k - i), snd (nth params i))) ps |
|
184 |
end; |
|
185 |
||
186 |
val prems = map (fn (prem, avoid) => |
|
187 |
let |
|
188 |
val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); |
|
189 |
val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); |
|
190 |
val params = Logic.strip_params prem |
|
191 |
in |
|
192 |
(params, |
|
193 |
fold (add_binders thy 0) (prems @ [concl]) [] @ |
|
194 |
map (apfst (incr_boundvars 1)) (mk_avoids params avoid), |
|
195 |
prems, strip_comb (HOLogic.dest_Trueprop concl)) |
|
196 |
end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); |
|
197 |
||
198 |
val atomTs = distinct op = (maps (map snd o #2) prems); |
|
199 |
val ind_sort = if null atomTs then HOLogic.typeS |
|
200 |
else Sign.certify_sort thy (map (fn T => Sign.intern_class thy |
|
201 |
("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs); |
|
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
202 |
val fs_ctxt_tyname = Name.variant (map fst (OldTerm.term_tfrees raw_induct')) "'n"; |
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
203 |
val fs_ctxt_name = Name.variant (OldTerm.add_term_names (raw_induct', [])) "z"; |
22530 | 204 |
val fsT = TFree (fs_ctxt_tyname, ind_sort); |
205 |
||
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
206 |
val inductive_forall_def' = Drule.instantiate' |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
207 |
[SOME (ctyp_of thy fsT)] [] inductive_forall_def; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
208 |
|
22530 | 209 |
fun lift_pred' t (Free (s, T)) ts = |
210 |
list_comb (Free (s, fsT --> T), t :: ts); |
|
211 |
val lift_pred = lift_pred' (Bound 0); |
|
212 |
||
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
213 |
fun lift_prem (t as (f $ u)) = |
22530 | 214 |
let val (p, ts) = strip_comb t |
215 |
in |
|
216 |
if p mem ps then |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
217 |
Const (inductive_forall_name, |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
218 |
(fsT --> HOLogic.boolT) --> HOLogic.boolT) $ |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
219 |
Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) |
22530 | 220 |
else lift_prem f $ lift_prem u |
221 |
end |
|
222 |
| lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) |
|
223 |
| lift_prem t = t; |
|
224 |
||
225 |
fun mk_distinct [] = [] |
|
226 |
| mk_distinct ((x, T) :: xs) = List.mapPartial (fn (y, U) => |
|
227 |
if T = U then SOME (HOLogic.mk_Trueprop |
|
228 |
(HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) |
|
229 |
else NONE) xs @ mk_distinct xs; |
|
230 |
||
231 |
fun mk_fresh (x, T) = HOLogic.mk_Trueprop |
|
25824 | 232 |
(NominalPackage.fresh_const T fsT $ x $ Bound 0); |
22530 | 233 |
|
234 |
val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => |
|
235 |
let |
|
236 |
val params' = params @ [("y", fsT)]; |
|
237 |
val prem = Logic.list_implies |
|
238 |
(map mk_fresh bvars @ mk_distinct bvars @ |
|
239 |
map (fn prem => |
|
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
27847
diff
changeset
|
240 |
if null (OldTerm.term_frees prem inter ps) then prem |
22530 | 241 |
else lift_prem prem) prems, |
242 |
HOLogic.mk_Trueprop (lift_pred p ts)); |
|
29276 | 243 |
val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') |
22530 | 244 |
in |
245 |
(list_all (params', prem), (rev vs, subst_bounds (vs, prem))) |
|
246 |
end) prems); |
|
247 |
||
248 |
val ind_vars = |
|
249 |
(DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ |
|
250 |
map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; |
|
251 |
val ind_Ts = rev (map snd ind_vars); |
|
252 |
||
253 |
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
254 |
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, |
|
255 |
HOLogic.list_all (ind_vars, lift_pred p |
|
256 |
(map (fold_rev (NominalPackage.mk_perm ind_Ts) |
|
257 |
(map Bound (length atomTs downto 1))) ts)))) concls)); |
|
258 |
||
259 |
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
260 |
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, |
|
261 |
lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); |
|
262 |
||
263 |
val vc_compat = map (fn (params, bvars, prems, (p, ts)) => |
|
264 |
map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
265 |
(List.mapPartial (fn prem => |
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
27847
diff
changeset
|
266 |
if null (ps inter OldTerm.term_frees prem) then SOME prem |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
267 |
else map_term (split_conj (K o I) names) prem prem) prems, q)))) |
22530 | 268 |
(mk_distinct bvars @ |
269 |
maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop |
|
25824 | 270 |
(NominalPackage.fresh_const U T $ u $ t)) bvars) |
22530 | 271 |
(ts ~~ binder_types (fastype_of p)))) prems; |
272 |
||
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
273 |
val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp"; |
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
274 |
val pt2_atoms = map (fn aT => PureThy.get_thm thy |
26337 | 275 |
("pt_" ^ Sign.base_name (fst (dest_Type aT)) ^ "2")) atomTs; |
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
276 |
val eqvt_ss = Simplifier.theory_context thy HOL_basic_ss |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
277 |
addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
278 |
addsimprocs [mk_perm_bool_simproc ["Fun.id"], |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
279 |
NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; |
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
280 |
val fresh_bij = PureThy.get_thms thy "fresh_bij"; |
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
281 |
val perm_bij = PureThy.get_thms thy "perm_bij"; |
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
282 |
val fs_atoms = map (fn aT => PureThy.get_thm thy |
26337 | 283 |
("fs_" ^ Sign.base_name (fst (dest_Type aT)) ^ "1")) atomTs; |
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
284 |
val exists_fresh' = PureThy.get_thms thy "exists_fresh'"; |
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
285 |
val fresh_atm = PureThy.get_thms thy "fresh_atm"; |
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
286 |
val swap_simps = PureThy.get_thms thy "swap_simps"; |
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
287 |
val perm_fresh_fresh = PureThy.get_thms thy "perm_fresh_fresh"; |
22530 | 288 |
|
289 |
fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = |
|
290 |
let |
|
27722
d657036e26c5
- corrected bogus comment for function inst_conj_all
berghofe
parents:
27449
diff
changeset
|
291 |
(** protect terms to avoid that fresh_prod interferes with **) |
22530 | 292 |
(** pairs used in introduction rules of inductive predicate **) |
293 |
fun protect t = |
|
294 |
let val T = fastype_of t in Const ("Fun.id", T --> T) $ t end; |
|
295 |
val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); |
|
296 |
val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop |
|
297 |
(HOLogic.exists_const T $ Abs ("x", T, |
|
25824 | 298 |
NominalPackage.fresh_const T (fastype_of p) $ |
22530 | 299 |
Bound 0 $ p))) |
300 |
(fn _ => EVERY |
|
301 |
[resolve_tac exists_fresh' 1, |
|
27722
d657036e26c5
- corrected bogus comment for function inst_conj_all
berghofe
parents:
27449
diff
changeset
|
302 |
resolve_tac fs_atoms 1]); |
22530 | 303 |
val (([cx], ths), ctxt') = Obtain.result |
304 |
(fn _ => EVERY |
|
305 |
[etac exE 1, |
|
306 |
full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, |
|
26359 | 307 |
full_simp_tac (HOL_basic_ss addsimps [@{thm id_apply}]) 1, |
22530 | 308 |
REPEAT (etac conjE 1)]) |
309 |
[ex] ctxt |
|
310 |
in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; |
|
311 |
||
25824 | 312 |
fun mk_ind_proof thy thss = |
26711 | 313 |
Goal.prove_global thy [] prems' concl' (fn {prems = ihyps, context = ctxt} => |
22530 | 314 |
let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => |
315 |
rtac raw_induct 1 THEN |
|
316 |
EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => |
|
317 |
[REPEAT (rtac allI 1), simp_tac eqvt_ss 1, |
|
318 |
SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => |
|
319 |
let |
|
320 |
val (params', (pis, z)) = |
|
321 |
chop (length params - length atomTs - 1) (map term_of params) ||> |
|
322 |
split_last; |
|
323 |
val bvars' = map |
|
324 |
(fn (Bound i, T) => (nth params' (length params' - i), T) |
|
325 |
| (t, T) => (t, T)) bvars; |
|
326 |
val pi_bvars = map (fn (t, _) => |
|
327 |
fold_rev (NominalPackage.mk_perm []) pis t) bvars'; |
|
328 |
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl)); |
|
329 |
val (freshs1, freshs2, ctxt'') = fold |
|
330 |
(obtain_fresh_name (ts @ pi_bvars)) |
|
331 |
(map snd bvars') ([], [], ctxt'); |
|
332 |
val freshs2' = NominalPackage.mk_not_sym freshs2; |
|
333 |
val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1); |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
334 |
fun concat_perm pi1 pi2 = |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
335 |
let val T = fastype_of pi1 |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
336 |
in if T = fastype_of pi2 then |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
337 |
Const ("List.append", T --> T --> T) $ pi1 $ pi2 |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
338 |
else pi2 |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
339 |
end; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
340 |
val pis'' = fold (concat_perm #> map) pis' pis; |
22530 | 341 |
val env = Pattern.first_order_match thy (ihypt, prop_of ihyp) |
342 |
(Vartab.empty, Vartab.empty); |
|
343 |
val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy)) |
|
344 |
(map (Envir.subst_vars env) vs ~~ |
|
345 |
map (fold_rev (NominalPackage.mk_perm []) |
|
346 |
(rev pis' @ pis)) params' @ [z])) ihyp; |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
347 |
fun mk_pi th = |
26359 | 348 |
Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}] |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
349 |
addsimprocs [NominalPackage.perm_simproc]) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
350 |
(Simplifier.simplify eqvt_ss |
25824 | 351 |
(fold_rev (mk_perm_bool o cterm_of thy) |
352 |
(rev pis' @ pis) th)); |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
353 |
val (gprems1, gprems2) = split_list |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
354 |
(map (fn (th, t) => |
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
27847
diff
changeset
|
355 |
if null (OldTerm.term_frees t inter ps) then (SOME th, mk_pi th) |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
356 |
else |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
357 |
(map_thm ctxt (split_conj (K o I) names) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
358 |
(etac conjunct1 1) monos NONE th, |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
359 |
mk_pi (the (map_thm ctxt (inst_conj_all names ps (rev pis'')) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
360 |
(inst_conj_all_tac (length pis'')) monos (SOME t) th)))) |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
361 |
(gprems ~~ oprems)) |>> List.mapPartial I; |
22530 | 362 |
val vc_compat_ths' = map (fn th => |
363 |
let |
|
25824 | 364 |
val th' = first_order_mrs gprems1 th; |
22530 | 365 |
val (bop, lhs, rhs) = (case concl_of th' of |
366 |
_ $ (fresh $ lhs $ rhs) => |
|
367 |
(fn t => fn u => fresh $ t $ u, lhs, rhs) |
|
368 |
| _ $ (_ $ (_ $ lhs $ rhs)) => |
|
369 |
(curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); |
|
370 |
val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop |
|
371 |
(bop (fold_rev (NominalPackage.mk_perm []) pis lhs) |
|
372 |
(fold_rev (NominalPackage.mk_perm []) pis rhs))) |
|
373 |
(fn _ => simp_tac (HOL_basic_ss addsimps |
|
374 |
(fresh_bij @ perm_bij)) 1 THEN rtac th' 1) |
|
375 |
in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end) |
|
376 |
vc_compat_ths; |
|
377 |
val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths'; |
|
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
378 |
(** 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:
27353
diff
changeset
|
379 |
(** we have to pre-simplify the rewrite rules **) |
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
380 |
val swap_simps_ss = HOL_ss addsimps swap_simps @ |
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
381 |
map (Simplifier.simplify (HOL_ss addsimps swap_simps)) |
22530 | 382 |
(vc_compat_ths'' @ freshs2'); |
383 |
val th = Goal.prove ctxt'' [] [] |
|
384 |
(HOLogic.mk_Trueprop (list_comb (P $ hd ts, |
|
385 |
map (fold (NominalPackage.mk_perm []) pis') (tl ts)))) |
|
386 |
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1, |
|
387 |
REPEAT_DETERM_N (nprems_of ihyp - length gprems) |
|
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
388 |
(simp_tac swap_simps_ss 1), |
22530 | 389 |
REPEAT_DETERM_N (length gprems) |
27847
0dffedf9aff5
Changed proof of strong induction rule to avoid infinite loop
berghofe
parents:
27722
diff
changeset
|
390 |
(simp_tac (HOL_basic_ss |
0dffedf9aff5
Changed proof of strong induction rule to avoid infinite loop
berghofe
parents:
27722
diff
changeset
|
391 |
addsimps [inductive_forall_def'] |
0dffedf9aff5
Changed proof of strong induction rule to avoid infinite loop
berghofe
parents:
27722
diff
changeset
|
392 |
addsimprocs [NominalPackage.perm_simproc]) 1 THEN |
0dffedf9aff5
Changed proof of strong induction rule to avoid infinite loop
berghofe
parents:
27722
diff
changeset
|
393 |
resolve_tac gprems2 1)])); |
22530 | 394 |
val final = Goal.prove ctxt'' [] [] (term_of concl) |
395 |
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss |
|
396 |
addsimps vc_compat_ths'' @ freshs2' @ |
|
397 |
perm_fresh_fresh @ fresh_atm) 1); |
|
398 |
val final' = ProofContext.export ctxt'' ctxt' [final]; |
|
399 |
in resolve_tac final' 1 end) context 1]) |
|
400 |
(prems ~~ thss ~~ ihyps ~~ prems''))) |
|
401 |
in |
|
402 |
cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN |
|
403 |
REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN |
|
27228
4f7976a6ffc3
allE_Nil: only one copy, proven in regular theory source;
wenzelm
parents:
27153
diff
changeset
|
404 |
etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN |
22530 | 405 |
asm_full_simp_tac (simpset_of thy) 1) |
26711 | 406 |
end); |
22530 | 407 |
|
25824 | 408 |
(** strong case analysis rule **) |
409 |
||
410 |
val cases_prems = map (fn ((name, avoids), rule) => |
|
411 |
let |
|
412 |
val prem :: prems = Logic.strip_imp_prems rule; |
|
413 |
val concl = Logic.strip_imp_concl rule; |
|
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
414 |
val used = OldTerm.add_term_free_names (rule, []) |
25824 | 415 |
in |
416 |
(prem, |
|
417 |
List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), |
|
418 |
concl, |
|
419 |
fst (fold_map (fn (prem, (_, avoid)) => fn used => |
|
420 |
let |
|
421 |
val prems = Logic.strip_assums_hyp prem; |
|
422 |
val params = Logic.strip_params prem; |
|
423 |
val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; |
|
424 |
fun mk_subst (p as (s, T)) (i, j, used, ps, qs, is, ts) = |
|
425 |
if member (op = o apsnd fst) bnds (Bound i) then |
|
426 |
let |
|
427 |
val s' = Name.variant used s; |
|
428 |
val t = Free (s', T) |
|
429 |
in (i + 1, j, s' :: used, ps, (t, T) :: qs, i :: is, t :: ts) end |
|
430 |
else (i + 1, j + 1, used, p :: ps, qs, is, Bound j :: ts); |
|
431 |
val (_, _, used', ps, qs, is, ts) = fold_rev mk_subst params |
|
432 |
(0, 0, used, [], [], [], []) |
|
433 |
in |
|
434 |
((ps, qs, is, map (curry subst_bounds (rev ts)) prems), used') |
|
435 |
end) (prems ~~ avoids) used)) |
|
436 |
end) |
|
437 |
(InductivePackage.partition_rules' raw_induct (intrs ~~ avoids') ~~ |
|
438 |
elims'); |
|
439 |
||
440 |
val cases_prems' = |
|
441 |
map (fn (prem, args, concl, prems) => |
|
442 |
let |
|
443 |
fun mk_prem (ps, [], _, prems) = |
|
444 |
list_all (ps, Logic.list_implies (prems, concl)) |
|
445 |
| mk_prem (ps, qs, _, prems) = |
|
446 |
list_all (ps, Logic.mk_implies |
|
447 |
(Logic.list_implies |
|
448 |
(mk_distinct qs @ |
|
449 |
maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop |
|
450 |
(NominalPackage.fresh_const T (fastype_of u) $ t $ u)) |
|
451 |
args) qs, |
|
452 |
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
453 |
(map HOLogic.dest_Trueprop prems))), |
|
454 |
concl)) |
|
455 |
in map mk_prem prems end) cases_prems; |
|
456 |
||
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
457 |
val cases_eqvt_ss = Simplifier.theory_context thy HOL_ss |
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
458 |
addsimps eqvt_thms @ swap_simps @ perm_pi_simp |
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
459 |
addsimprocs [NominalPermeq.perm_simproc_app, |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
460 |
NominalPermeq.perm_simproc_fun]; |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
461 |
|
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
462 |
val simp_fresh_atm = map |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
463 |
(Simplifier.simplify (HOL_basic_ss addsimps fresh_atm)); |
25824 | 464 |
|
465 |
fun mk_cases_proof thy ((((name, thss), elim), (prem, args, concl, prems)), |
|
466 |
prems') = |
|
26711 | 467 |
(name, Goal.prove_global thy [] (prem :: prems') concl |
468 |
(fn {prems = hyp :: hyps, context = ctxt1} => |
|
25824 | 469 |
EVERY (rtac (hyp RS elim) 1 :: |
470 |
map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => |
|
471 |
SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} => |
|
472 |
if null qs then |
|
473 |
rtac (first_order_mrs case_hyps case_hyp) 1 |
|
474 |
else |
|
475 |
let |
|
476 |
val params' = map (term_of o nth (rev params)) is; |
|
477 |
val tab = params' ~~ map fst qs; |
|
478 |
val (hyps1, hyps2) = chop (length args) case_hyps; |
|
479 |
(* turns a = t and [x1 # t, ..., xn # t] *) |
|
480 |
(* into [x1 # a, ..., xn # a] *) |
|
481 |
fun inst_fresh th' ths = |
|
482 |
let val (ths1, ths2) = chop (length qs) ths |
|
483 |
in |
|
484 |
(map (fn th => |
|
485 |
let |
|
486 |
val (cf, ct) = |
|
487 |
Thm.dest_comb (Thm.dest_arg (cprop_of th)); |
|
488 |
val arg_cong' = Drule.instantiate' |
|
489 |
[SOME (ctyp_of_term ct)] |
|
490 |
[NONE, SOME ct, SOME cf] (arg_cong RS iffD2); |
|
491 |
val inst = Thm.first_order_match (ct, |
|
492 |
Thm.dest_arg (Thm.dest_arg (cprop_of th'))) |
|
493 |
in [th', th] MRS Thm.instantiate inst arg_cong' |
|
494 |
end) ths1, |
|
495 |
ths2) |
|
496 |
end; |
|
497 |
val (vc_compat_ths1, vc_compat_ths2) = |
|
498 |
chop (length vc_compat_ths - length args * length qs) |
|
499 |
(map (first_order_mrs hyps2) vc_compat_ths); |
|
500 |
val vc_compat_ths' = |
|
501 |
NominalPackage.mk_not_sym vc_compat_ths1 @ |
|
502 |
flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); |
|
503 |
val (freshs1, freshs2, ctxt3) = fold |
|
504 |
(obtain_fresh_name (args @ map fst qs @ params')) |
|
505 |
(map snd qs) ([], [], ctxt2); |
|
506 |
val freshs2' = NominalPackage.mk_not_sym freshs2; |
|
507 |
val pis = map (NominalPackage.perm_of_pair) |
|
508 |
((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); |
|
509 |
val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis); |
|
510 |
val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms |
|
511 |
(fn x as Free _ => |
|
512 |
if x mem args then x |
|
513 |
else (case AList.lookup op = tab x of |
|
514 |
SOME y => y |
|
515 |
| NONE => fold_rev (NominalPackage.mk_perm []) pis x) |
|
516 |
| x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps)); |
|
517 |
val inst = Thm.first_order_match (Thm.dest_arg |
|
518 |
(Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); |
|
519 |
val th = Goal.prove ctxt3 [] [] (term_of concl) |
|
520 |
(fn {context = ctxt4, ...} => |
|
521 |
rtac (Thm.instantiate inst case_hyp) 1 THEN |
|
522 |
SUBPROOF (fn {prems = fresh_hyps, ...} => |
|
523 |
let |
|
524 |
val fresh_hyps' = NominalPackage.mk_not_sym fresh_hyps; |
|
27352
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
525 |
val case_ss = cases_eqvt_ss addsimps freshs2' @ |
8adeff7fd4bc
- Equivariance simpset used in proofs of strong induction and inversion
berghofe
parents:
27228
diff
changeset
|
526 |
simp_fresh_atm (vc_compat_ths' @ fresh_hyps'); |
25824 | 527 |
val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh; |
528 |
val hyps1' = map |
|
529 |
(mk_pis #> Simplifier.simplify fresh_fresh_ss) hyps1; |
|
530 |
val hyps2' = map |
|
531 |
(mk_pis #> Simplifier.simplify case_ss) hyps2; |
|
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
532 |
val case_hyps' = hyps1' @ hyps2' |
25824 | 533 |
in |
27449
4880da911af0
Rewrote code to use swap_simps rather than calc_atm (which tends to
berghofe
parents:
27353
diff
changeset
|
534 |
simp_tac case_ss 1 THEN |
25824 | 535 |
REPEAT_DETERM (TRY (rtac conjI 1) THEN |
536 |
resolve_tac case_hyps' 1) |
|
537 |
end) ctxt4 1) |
|
538 |
val final = ProofContext.export ctxt3 ctxt2 [th] |
|
539 |
in resolve_tac final 1 end) ctxt1 1) |
|
540 |
(thss ~~ hyps ~~ prems)))) |
|
541 |
||
22530 | 542 |
in |
543 |
thy |> |
|
544 |
ProofContext.init |> |
|
545 |
Proof.theorem_i NONE (fn thss => ProofContext.theory (fn thy => |
|
546 |
let |
|
547 |
val ctxt = ProofContext.init thy; |
|
548 |
val rec_name = space_implode "_" (map Sign.base_name names); |
|
549 |
val ind_case_names = RuleCases.case_names induct_cases; |
|
25824 | 550 |
val induct_cases' = InductivePackage.partition_rules' raw_induct |
551 |
(intrs ~~ induct_cases); |
|
552 |
val thss' = map (map atomize_intr) thss; |
|
553 |
val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss'); |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
554 |
val strong_raw_induct = |
25824 | 555 |
mk_ind_proof thy thss' |> InductivePackage.rulify; |
556 |
val strong_cases = map (mk_cases_proof thy ##> InductivePackage.rulify) |
|
557 |
(thsss ~~ elims ~~ cases_prems ~~ cases_prems'); |
|
22530 | 558 |
val strong_induct = |
559 |
if length names > 1 then |
|
560 |
(strong_raw_induct, [ind_case_names, RuleCases.consumes 0]) |
|
561 |
else (strong_raw_induct RSN (2, rev_mp), |
|
562 |
[ind_case_names, RuleCases.consumes 1]); |
|
563 |
val ([strong_induct'], thy') = thy |> |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24571
diff
changeset
|
564 |
Sign.add_path rec_name |> |
22530 | 565 |
PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)]; |
566 |
val strong_inducts = |
|
567 |
ProjectRule.projects ctxt (1 upto length names) strong_induct' |
|
568 |
in |
|
569 |
thy' |> |
|
570 |
PureThy.add_thmss [(("strong_inducts", strong_inducts), |
|
571 |
[ind_case_names, RuleCases.consumes 1])] |> snd |> |
|
25824 | 572 |
Sign.parent_path |> |
573 |
fold (fn ((name, elim), (_, cases)) => |
|
574 |
Sign.add_path (Sign.base_name name) #> |
|
575 |
PureThy.add_thms [(("strong_cases", elim), |
|
576 |
[RuleCases.case_names (map snd cases), |
|
577 |
RuleCases.consumes 1])] #> snd #> |
|
578 |
Sign.parent_path) (strong_cases ~~ induct_cases') |
|
22530 | 579 |
end)) |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
580 |
(map (map (rulify_term thy #> rpair [])) vc_compat) |
22530 | 581 |
end; |
582 |
||
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
583 |
fun prove_eqvt s xatoms thy = |
22530 | 584 |
let |
585 |
val ctxt = ProofContext.init thy; |
|
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
586 |
val ({names, ...}, {raw_induct, intrs, elims, ...}) = |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
587 |
InductivePackage.the_inductive ctxt (Sign.intern_const thy s); |
24832 | 588 |
val raw_induct = atomize_induct ctxt raw_induct; |
589 |
val elims = map (atomize_induct ctxt) elims; |
|
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
590 |
val intrs = map atomize_intr intrs; |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
591 |
val monos = InductivePackage.get_monos ctxt; |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
592 |
val intrs' = InductivePackage.unpartition_rules intrs |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
593 |
(map (fn (((s, ths), (_, k)), th) => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
594 |
(s, ths ~~ InductivePackage.infer_intro_vars th k ths)) |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
595 |
(InductivePackage.partition_rules raw_induct intrs ~~ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
596 |
InductivePackage.arities_of raw_induct ~~ elims)); |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
597 |
val atoms' = NominalAtoms.atoms_of thy; |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
598 |
val atoms = |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
599 |
if null xatoms then atoms' else |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
600 |
let val atoms = map (Sign.intern_type thy) xatoms |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
601 |
in |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
602 |
(case duplicates op = atoms of |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
603 |
[] => () |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
604 |
| xs => error ("Duplicate atoms: " ^ commas xs); |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
605 |
case atoms \\ atoms' of |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
606 |
[] => () |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
607 |
| xs => error ("No such atoms: " ^ commas xs); |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
608 |
atoms) |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
609 |
end; |
26343
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
wenzelm
parents:
26337
diff
changeset
|
610 |
val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp"; |
26364
cb6f360ab425
Equivariance prover now uses permutation simprocs as well.
berghofe
parents:
26359
diff
changeset
|
611 |
val eqvt_ss = Simplifier.theory_context thy HOL_basic_ss addsimps |
24571 | 612 |
(NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs |
26364
cb6f360ab425
Equivariance prover now uses permutation simprocs as well.
berghofe
parents:
26359
diff
changeset
|
613 |
[mk_perm_bool_simproc names, |
cb6f360ab425
Equivariance prover now uses permutation simprocs as well.
berghofe
parents:
26359
diff
changeset
|
614 |
NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
615 |
val t = Logic.unvarify (concl_of raw_induct); |
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
616 |
val pi = Name.variant (OldTerm.add_term_names (t, [])) "pi"; |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
617 |
val ps = map (fst o HOLogic.dest_imp) |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
618 |
(HOLogic.dest_conj (HOLogic.dest_Trueprop t)); |
25824 | 619 |
fun eqvt_tac pi (intr, vs) st = |
22544 | 620 |
let |
621 |
fun eqvt_err s = error |
|
622 |
("Could not prove equivariance for introduction rule\n" ^ |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26711
diff
changeset
|
623 |
Syntax.string_of_term_global (theory_of_thm intr) |
22544 | 624 |
(Logic.unvarify (prop_of intr)) ^ "\n" ^ s); |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
625 |
val res = SUBPROOF (fn {prems, params, ...} => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
626 |
let |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
627 |
val prems' = map (fn th => the_default th (map_thm ctxt |
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
628 |
(split_conj (K I) names) (etac conjunct2 1) monos NONE th)) prems; |
25824 | 629 |
val prems'' = map (fn th => Simplifier.simplify eqvt_ss |
630 |
(mk_perm_bool (cterm_of thy pi) th)) prems'; |
|
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
631 |
val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
632 |
map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params) |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
633 |
intr |
24570
621b60b1df00
Generalized equivariance and nominal_inductive commands to
berghofe
parents:
23531
diff
changeset
|
634 |
in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1 |
22544 | 635 |
end) ctxt 1 st |
636 |
in |
|
637 |
case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of |
|
638 |
NONE => eqvt_err ("Rule does not match goal\n" ^ |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26711
diff
changeset
|
639 |
Syntax.string_of_term_global (theory_of_thm st) (hd (prems_of st))) |
22544 | 640 |
| SOME (th, _) => Seq.single th |
641 |
end; |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
642 |
val thss = map (fn atom => |
25824 | 643 |
let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) |
22530 | 644 |
in map (fn th => zero_var_indexes (th RS mp)) |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
645 |
(DatatypeAux.split_conj_thm (Goal.prove_global thy [] [] |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
646 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
647 |
HOLogic.mk_imp (p, list_comb |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
648 |
(apsnd (map (NominalPackage.mk_perm [] pi')) (strip_comb p)))) ps))) |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
649 |
(fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
650 |
full_simp_tac eqvt_ss 1 THEN |
25824 | 651 |
eqvt_tac pi' intr_vs) intrs')))) |
22544 | 652 |
end) atoms |
653 |
in |
|
654 |
fold (fn (name, ths) => |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24571
diff
changeset
|
655 |
Sign.add_path (Sign.base_name name) #> |
22544 | 656 |
PureThy.add_thmss [(("eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #> |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24571
diff
changeset
|
657 |
Sign.parent_path) (names ~~ transp thss) thy |
22544 | 658 |
end; |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
659 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
660 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
661 |
(* outer syntax *) |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
662 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
663 |
local structure P = OuterParse and K = OuterKeyword in |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
664 |
|
27353
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents:
27352
diff
changeset
|
665 |
val _ = OuterKeyword.keyword "avoids"; |
24867 | 666 |
|
667 |
val _ = |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
668 |
OuterSyntax.command "nominal_inductive" |
22530 | 669 |
"prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal |
670 |
(P.name -- Scan.optional (P.$$$ "avoids" |-- P.and_list1 (P.name -- |
|
671 |
(P.$$$ ":" |-- Scan.repeat1 P.name))) [] >> (fn (name, avoids) => |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
672 |
Toplevel.print o Toplevel.theory_to_proof (prove_strong_ind name avoids))); |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
673 |
|
24867 | 674 |
val _ = |
22530 | 675 |
OuterSyntax.command "equivariance" |
676 |
"prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
677 |
(P.name -- Scan.optional (P.$$$ "[" |-- P.list1 P.name --| P.$$$ "]") [] >> |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
678 |
(fn (name, atoms) => Toplevel.theory (prove_eqvt name atoms))); |
22530 | 679 |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
680 |
end; |
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 |
end |