author | wenzelm |
Tue, 11 Jul 2006 12:16:54 +0200 | |
changeset 20071 | 8f3e1ddb50e6 |
parent 20047 | e23a3afaaa8a |
child 20248 | 7916ce5bb069 |
permissions | -rw-r--r-- |
5094 | 1 |
(* Title: HOL/Tools/inductive_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
10735 | 4 |
Author: Stefan Berghofer, TU Muenchen |
5 |
Author: Markus Wenzel, TU Muenchen |
|
5094 | 6 |
|
6424 | 7 |
(Co)Inductive Definition module for HOL. |
5094 | 8 |
|
9 |
Features: |
|
6424 | 10 |
* least or greatest fixedpoints |
11 |
* user-specified product and sum constructions |
|
12 |
* mutually recursive definitions |
|
13 |
* definitions involving arbitrary monotone operators |
|
14 |
* automatically proves introduction and elimination rules |
|
5094 | 15 |
|
6424 | 16 |
The recursive sets must *already* be declared as constants in the |
17 |
current theory! |
|
5094 | 18 |
|
19 |
Introduction rules have the form |
|
8316
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
20 |
[| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |
5094 | 21 |
where M is some monotone operator (usually the identity) |
22 |
P(x) is any side condition on the free variables |
|
23 |
ti, t are any terms |
|
24 |
Sj, Sk are two of the sets being defined in mutual recursion |
|
25 |
||
6424 | 26 |
Sums are used only for mutual recursion. Products are used only to |
27 |
derive "streamlined" induction rules for relations. |
|
5094 | 28 |
*) |
29 |
||
30 |
signature INDUCTIVE_PACKAGE = |
|
31 |
sig |
|
6424 | 32 |
val quiet_mode: bool ref |
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
33 |
val trace: bool ref |
16432 | 34 |
val unify_consts: theory -> term list -> term list -> term list * term list |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
35 |
val split_rule_vars: term list -> thm -> thm |
9116
9df44b5c610b
get_inductive now returns None instead of raising an exception.
berghofe
parents:
9072
diff
changeset
|
36 |
val get_inductive: theory -> string -> ({names: string list, coind: bool} * |
9df44b5c610b
get_inductive now returns None instead of raising an exception.
berghofe
parents:
9072
diff
changeset
|
37 |
{defs: thm list, elims: thm list, raw_induct: thm, induct: thm, |
9df44b5c610b
get_inductive now returns None instead of raising an exception.
berghofe
parents:
9072
diff
changeset
|
38 |
intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}) option |
12400 | 39 |
val the_mk_cases: theory -> string -> string -> thm |
6437 | 40 |
val print_inductives: theory -> unit |
18728 | 41 |
val mono_add: attribute |
42 |
val mono_del: attribute |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
43 |
val get_monos: theory -> thm list |
10910
058775a575db
export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents:
10804
diff
changeset
|
44 |
val inductive_forall_name: string |
058775a575db
export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents:
10804
diff
changeset
|
45 |
val inductive_forall_def: thm |
058775a575db
export inductive_forall_name, inductive_forall_def, rulify;
wenzelm
parents:
10804
diff
changeset
|
46 |
val rulify: thm -> thm |
15703 | 47 |
val inductive_cases: ((bstring * Attrib.src list) * string list) list -> theory -> theory |
18728 | 48 |
val inductive_cases_i: ((bstring * attribute list) * term list) list -> theory -> theory |
6424 | 49 |
val add_inductive_i: bool -> bool -> bstring -> bool -> bool -> bool -> term list -> |
18728 | 50 |
((bstring * term) * attribute list) list -> thm list -> theory -> theory * |
6424 | 51 |
{defs: thm list, elims: thm list, raw_induct: thm, induct: thm, |
6437 | 52 |
intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm} |
11628 | 53 |
val add_inductive: bool -> bool -> string list -> |
15703 | 54 |
((bstring * string) * Attrib.src list) list -> (thmref * Attrib.src list) list -> |
12180 | 55 |
theory -> theory * |
6424 | 56 |
{defs: thm list, elims: thm list, raw_induct: thm, induct: thm, |
6437 | 57 |
intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm} |
18708 | 58 |
val setup: theory -> theory |
5094 | 59 |
end; |
60 |
||
6424 | 61 |
structure InductivePackage: INDUCTIVE_PACKAGE = |
5094 | 62 |
struct |
63 |
||
9598 | 64 |
|
10729 | 65 |
(** theory context references **) |
66 |
||
15525 | 67 |
val mono_name = "Orderings.mono"; |
17010
5abc26872268
changed reference to Lfp.lfp to FixedPint.lfp, ditto for gfp
avigad
parents:
16975
diff
changeset
|
68 |
val gfp_name = "FixedPoint.gfp"; |
5abc26872268
changed reference to Lfp.lfp to FixedPint.lfp, ditto for gfp
avigad
parents:
16975
diff
changeset
|
69 |
val lfp_name = "FixedPoint.lfp"; |
12259 | 70 |
val vimage_name = "Set.vimage"; |
10735 | 71 |
val Const _ $ (vimage_f $ _) $ _ = HOLogic.dest_Trueprop (Thm.concl_of vimageD); |
72 |
||
11991 | 73 |
val inductive_forall_name = "HOL.induct_forall"; |
74 |
val inductive_forall_def = thm "induct_forall_def"; |
|
75 |
val inductive_conj_name = "HOL.induct_conj"; |
|
76 |
val inductive_conj_def = thm "induct_conj_def"; |
|
77 |
val inductive_conj = thms "induct_conj"; |
|
78 |
val inductive_atomize = thms "induct_atomize"; |
|
18463 | 79 |
val inductive_rulify = thms "induct_rulify"; |
80 |
val inductive_rulify_fallback = thms "induct_rulify_fallback"; |
|
10729 | 81 |
|
82 |
||
83 |
||
10735 | 84 |
(** theory data **) |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
85 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
86 |
type inductive_info = |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
87 |
{names: string list, coind: bool} * {defs: thm list, elims: thm list, raw_induct: thm, |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
88 |
induct: thm, intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}; |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
89 |
|
16432 | 90 |
structure InductiveData = TheoryDataFun |
91 |
(struct |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
92 |
val name = "HOL/inductive"; |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
93 |
type T = inductive_info Symtab.table * thm list; |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
94 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
95 |
val empty = (Symtab.empty, []); |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
96 |
val copy = I; |
16432 | 97 |
val extend = I; |
98 |
fun merge _ ((tab1, monos1), (tab2, monos2)) = |
|
11502 | 99 |
(Symtab.merge (K true) (tab1, tab2), Drule.merge_rules (monos1, monos2)); |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
100 |
|
16432 | 101 |
fun print thy (tab, monos) = |
16364 | 102 |
[Pretty.strs ("(co)inductives:" :: |
16432 | 103 |
map #1 (NameSpace.extern_table (Sign.const_space thy, tab))), |
104 |
Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm_sg thy) monos)] |
|
8720 | 105 |
|> Pretty.chunks |> Pretty.writeln; |
16432 | 106 |
end); |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
107 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
108 |
val print_inductives = InductiveData.print; |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
109 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
110 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
111 |
(* get and put data *) |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
112 |
|
17412 | 113 |
val get_inductive = Symtab.lookup o #1 o InductiveData.get; |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
114 |
|
9598 | 115 |
fun the_inductive thy name = |
116 |
(case get_inductive thy name of |
|
15531 | 117 |
NONE => error ("Unknown (co)inductive set " ^ quote name) |
118 |
| SOME info => info); |
|
9598 | 119 |
|
12400 | 120 |
val the_mk_cases = (#mk_cases o #2) oo the_inductive; |
121 |
||
18222 | 122 |
fun put_inductives names info = InductiveData.map (apfst (fn tab => |
123 |
fold (fn name => Symtab.update_new (name, info)) names tab |
|
124 |
handle Symtab.DUP dup => error ("Duplicate definition of (co)inductive set " ^ quote dup))); |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
125 |
|
8277 | 126 |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
127 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
128 |
(** monotonicity rules **) |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
129 |
|
9831 | 130 |
val get_monos = #2 o InductiveData.get; |
18728 | 131 |
val map_monos = Context.map_theory o InductiveData.map o Library.apsnd; |
8277 | 132 |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
133 |
fun mk_mono thm = |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
134 |
let |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
135 |
fun eq2mono thm' = [standard (thm' RS (thm' RS eq_to_mono))] @ |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
136 |
(case concl_of thm of |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
137 |
(_ $ (_ $ (Const ("Not", _) $ _) $ _)) => [] |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
138 |
| _ => [standard (thm' RS (thm' RS eq_to_mono2))]); |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
139 |
val concl = concl_of thm |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
140 |
in |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
141 |
if Logic.is_equals concl then |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
142 |
eq2mono (thm RS meta_eq_to_obj_eq) |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
143 |
else if can (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl then |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
144 |
eq2mono thm |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
145 |
else [thm] |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
146 |
end; |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
147 |
|
8634 | 148 |
|
149 |
(* attributes *) |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
150 |
|
18921 | 151 |
val mono_add = Thm.declaration_attribute (map_monos o fold Drule.add_rule o mk_mono); |
152 |
val mono_del = Thm.declaration_attribute (map_monos o fold Drule.del_rule o mk_mono); |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
153 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
154 |
|
7107 | 155 |
|
10735 | 156 |
(** misc utilities **) |
6424 | 157 |
|
5662 | 158 |
val quiet_mode = ref false; |
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
159 |
val trace = ref false; (*for debugging*) |
10735 | 160 |
fun message s = if ! quiet_mode then () else writeln s; |
161 |
fun clean_message s = if ! quick_and_dirty then () else message s; |
|
5662 | 162 |
|
6424 | 163 |
fun coind_prefix true = "co" |
164 |
| coind_prefix false = ""; |
|
165 |
||
166 |
||
10735 | 167 |
(*the following code ensures that each recursive set always has the |
168 |
same type in all introduction rules*) |
|
16432 | 169 |
fun unify_consts thy cs intr_ts = |
7020
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
170 |
(let |
16861 | 171 |
val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I); |
7020
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
172 |
fun varify (t, (i, ts)) = |
16876 | 173 |
let val t' = map_term_types (Logic.incr_tvar (i + 1)) (#1 (Type.varify (t, []))) |
7020
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
174 |
in (maxidx_of_term t', t'::ts) end; |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
175 |
val (i, cs') = foldr varify (~1, []) cs; |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
176 |
val (i', intr_ts') = foldr varify (i, []) intr_ts; |
16861 | 177 |
val rec_consts = fold add_term_consts_2 cs' []; |
178 |
val intr_consts = fold add_term_consts_2 intr_ts' []; |
|
16934 | 179 |
fun unify (cname, cT) = |
15570 | 180 |
let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts) |
16934 | 181 |
in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end; |
182 |
val (env, _) = fold unify rec_consts (Vartab.empty, i'); |
|
16287 | 183 |
val subst = Type.freeze o map_term_types (Envir.norm_type env) |
7020
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
184 |
|
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
185 |
in (map subst cs', map subst intr_ts') |
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
186 |
end) handle Type.TUNIFY => |
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
187 |
(warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)); |
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
188 |
|
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
189 |
|
10735 | 190 |
(*make injections used in mutually recursive definitions*) |
5094 | 191 |
fun mk_inj cs sumT c x = |
192 |
let |
|
193 |
fun mk_inj' T n i = |
|
194 |
if n = 1 then x else |
|
195 |
let val n2 = n div 2; |
|
196 |
val Type (_, [T1, T2]) = T |
|
197 |
in |
|
198 |
if i <= n2 then |
|
15391
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents:
15032
diff
changeset
|
199 |
Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i) |
5094 | 200 |
else |
15391
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents:
15032
diff
changeset
|
201 |
Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) |
5094 | 202 |
end |
203 |
in mk_inj' sumT (length cs) (1 + find_index_eq c cs) |
|
204 |
end; |
|
205 |
||
10735 | 206 |
(*make "vimage" terms for selecting out components of mutually rec.def*) |
5094 | 207 |
fun mk_vimage cs sumT t c = if length cs < 2 then t else |
208 |
let |
|
209 |
val cT = HOLogic.dest_setT (fastype_of c); |
|
210 |
val vimageT = [cT --> sumT, HOLogic.mk_setT sumT] ---> HOLogic.mk_setT cT |
|
211 |
in |
|
212 |
Const (vimage_name, vimageT) $ |
|
213 |
Abs ("y", cT, mk_inj cs sumT c (Bound 0)) $ t |
|
214 |
end; |
|
215 |
||
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
216 |
(** proper splitting **) |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
217 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
218 |
fun prod_factors p (Const ("Pair", _) $ t $ u) = |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
219 |
p :: prod_factors (1::p) t @ prod_factors (2::p) u |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
220 |
| prod_factors p _ = []; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
221 |
|
17485 | 222 |
fun mg_prod_factors ts (t $ u) fs = if t mem ts then |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
223 |
let val f = prod_factors [] u |
17325 | 224 |
in AList.update (op =) (t, f inter (AList.lookup (op =) fs t) |> the_default f) fs end |
17485 | 225 |
else mg_prod_factors ts u (mg_prod_factors ts t fs) |
226 |
| mg_prod_factors ts (Abs (_, _, t)) fs = mg_prod_factors ts t fs |
|
227 |
| mg_prod_factors ts _ fs = fs; |
|
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
228 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
229 |
fun prodT_factors p ps (T as Type ("*", [T1, T2])) = |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
230 |
if p mem ps then prodT_factors (1::p) ps T1 @ prodT_factors (2::p) ps T2 |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
231 |
else [T] |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
232 |
| prodT_factors _ _ T = [T]; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
233 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
234 |
fun ap_split p ps (Type ("*", [T1, T2])) T3 u = |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
235 |
if p mem ps then HOLogic.split_const (T1, T2, T3) $ |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
236 |
Abs ("v", T1, ap_split (2::p) ps T2 T3 (ap_split (1::p) ps T1 |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
237 |
(prodT_factors (2::p) ps T2 ---> T3) (incr_boundvars 1 u) $ Bound 0)) |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
238 |
else u |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
239 |
| ap_split _ _ _ _ u = u; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
240 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
241 |
fun mk_tuple p ps (Type ("*", [T1, T2])) (tms as t::_) = |
17985 | 242 |
if p mem ps then HOLogic.mk_prod (mk_tuple (1::p) ps T1 tms, |
15570 | 243 |
mk_tuple (2::p) ps T2 (Library.drop (length (prodT_factors (1::p) ps T1), tms))) |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
244 |
else t |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
245 |
| mk_tuple _ _ _ (t::_) = t; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
246 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
247 |
fun split_rule_var' ((t as Var (v, Type ("fun", [T1, T2])), ps), rl) = |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
248 |
let val T' = prodT_factors [] ps T1 ---> T2 |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
249 |
val newt = ap_split [] ps T1 T2 (Var (v, T')) |
16432 | 250 |
val cterm = Thm.cterm_of (Thm.theory_of_thm rl) |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
251 |
in |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
252 |
instantiate ([], [(cterm t, cterm newt)]) rl |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
253 |
end |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
254 |
| split_rule_var' (_, rl) = rl; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
255 |
|
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
256 |
val remove_split = rewrite_rule [split_conv RS eq_reflection]; |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
257 |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
258 |
fun split_rule_vars vs rl = standard (remove_split (foldr split_rule_var' |
17485 | 259 |
rl (mg_prod_factors vs (Thm.prop_of rl) []))); |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
260 |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
261 |
fun split_rule vs rl = standard (remove_split (foldr split_rule_var' |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
262 |
rl (List.mapPartial (fn (t as Var ((a, _), _)) => |
17485 | 263 |
Option.map (pair t) (AList.lookup (op =) vs a)) (term_vars (Thm.prop_of rl))))); |
6424 | 264 |
|
265 |
||
10729 | 266 |
(** process rules **) |
267 |
||
268 |
local |
|
5094 | 269 |
|
16432 | 270 |
fun err_in_rule thy name t msg = |
271 |
error (cat_lines ["Ill-formed introduction rule " ^ quote name, |
|
272 |
Sign.string_of_term thy t, msg]); |
|
10729 | 273 |
|
16432 | 274 |
fun err_in_prem thy name t p msg = |
275 |
error (cat_lines ["Ill-formed premise", Sign.string_of_term thy p, |
|
276 |
"in introduction rule " ^ quote name, Sign.string_of_term thy t, msg]); |
|
5094 | 277 |
|
10729 | 278 |
val bad_concl = "Conclusion of introduction rule must have form \"t : S_i\""; |
279 |
||
17985 | 280 |
val all_not_allowed = |
11358
416ea5c009f5
now checks for leading meta-quantifiers and complains, instead of
paulson
parents:
11036
diff
changeset
|
281 |
"Introduction rule must not have a leading \"!!\" quantifier"; |
416ea5c009f5
now checks for leading meta-quantifiers and complains, instead of
paulson
parents:
11036
diff
changeset
|
282 |
|
16432 | 283 |
fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize []; |
10729 | 284 |
|
285 |
in |
|
5094 | 286 |
|
16432 | 287 |
fun check_rule thy cs ((name, rule), att) = |
10729 | 288 |
let |
289 |
val concl = Logic.strip_imp_concl rule; |
|
290 |
val prems = Logic.strip_imp_prems rule; |
|
16432 | 291 |
val aprems = map (atomize_term thy) prems; |
10729 | 292 |
val arule = Logic.list_implies (aprems, concl); |
5094 | 293 |
|
10729 | 294 |
fun check_prem (prem, aprem) = |
295 |
if can HOLogic.dest_Trueprop aprem then () |
|
16432 | 296 |
else err_in_prem thy name rule prem "Non-atomic premise"; |
10729 | 297 |
in |
11358
416ea5c009f5
now checks for leading meta-quantifiers and complains, instead of
paulson
parents:
11036
diff
changeset
|
298 |
(case concl of |
416ea5c009f5
now checks for leading meta-quantifiers and complains, instead of
paulson
parents:
11036
diff
changeset
|
299 |
Const ("Trueprop", _) $ (Const ("op :", _) $ t $ u) => |
10729 | 300 |
if u mem cs then |
301 |
if exists (Logic.occs o rpair t) cs then |
|
16432 | 302 |
err_in_rule thy name rule "Recursion term on left of member symbol" |
15570 | 303 |
else List.app check_prem (prems ~~ aprems) |
16432 | 304 |
else err_in_rule thy name rule bad_concl |
305 |
| Const ("all", _) $ _ => err_in_rule thy name rule all_not_allowed |
|
306 |
| _ => err_in_rule thy name rule bad_concl); |
|
10729 | 307 |
((name, arule), att) |
308 |
end; |
|
5094 | 309 |
|
18222 | 310 |
val rulify = (* FIXME norm_hhf *) |
311 |
hol_simplify inductive_conj |
|
18463 | 312 |
#> hol_simplify inductive_rulify |
313 |
#> hol_simplify inductive_rulify_fallback |
|
18222 | 314 |
#> standard; |
10729 | 315 |
|
316 |
end; |
|
317 |
||
5094 | 318 |
|
6424 | 319 |
|
10735 | 320 |
(** properties of (co)inductive sets **) |
5094 | 321 |
|
10735 | 322 |
(* elimination rules *) |
5094 | 323 |
|
8375 | 324 |
fun mk_elims cs cTs params intr_ts intr_names = |
5094 | 325 |
let |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
326 |
val used = foldr add_term_names [] intr_ts; |
20071
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20047
diff
changeset
|
327 |
val [aname, pname] = Name.variant_list used ["a", "P"]; |
5094 | 328 |
val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT)); |
329 |
||
330 |
fun dest_intr r = |
|
331 |
let val Const ("op :", _) $ t $ u = |
|
332 |
HOLogic.dest_Trueprop (Logic.strip_imp_concl r) |
|
333 |
in (u, t, Logic.strip_imp_prems r) end; |
|
334 |
||
8380 | 335 |
val intrs = map dest_intr intr_ts ~~ intr_names; |
5094 | 336 |
|
337 |
fun mk_elim (c, T) = |
|
338 |
let |
|
339 |
val a = Free (aname, T); |
|
340 |
||
341 |
fun mk_elim_prem (_, t, ts) = |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
342 |
list_all_free (map dest_Free ((foldr add_term_frees [] (t::ts)) \\ params), |
5094 | 343 |
Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (a, t)) :: ts, P)); |
15570 | 344 |
val c_intrs = (List.filter (equal c o #1 o #1) intrs); |
5094 | 345 |
in |
8375 | 346 |
(Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (a, c)) :: |
347 |
map mk_elim_prem (map #1 c_intrs), P), map #2 c_intrs) |
|
5094 | 348 |
end |
349 |
in |
|
350 |
map mk_elim (cs ~~ cTs) |
|
351 |
end; |
|
9598 | 352 |
|
6424 | 353 |
|
10735 | 354 |
(* premises and conclusions of induction rules *) |
5094 | 355 |
|
356 |
fun mk_indrule cs cTs params intr_ts = |
|
357 |
let |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
358 |
val used = foldr add_term_names [] intr_ts; |
5094 | 359 |
|
360 |
(* predicates for induction rule *) |
|
361 |
||
20071
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20047
diff
changeset
|
362 |
val preds = map Free (Name.variant_list used (if length cs < 2 then ["P"] else |
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20047
diff
changeset
|
363 |
map (fn i => "P" ^ string_of_int i) (1 upto length cs)) ~~ |
5094 | 364 |
map (fn T => T --> HOLogic.boolT) cTs); |
365 |
||
366 |
(* transform an introduction rule into a premise for induction rule *) |
|
367 |
||
368 |
fun mk_ind_prem r = |
|
369 |
let |
|
370 |
val frees = map dest_Free ((add_term_frees (r, [])) \\ params); |
|
371 |
||
17485 | 372 |
val pred_of = AList.lookup (op aconv) (cs ~~ preds); |
5094 | 373 |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
374 |
fun subst (s as ((m as Const ("op :", T)) $ t $ u)) = |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
375 |
(case pred_of u of |
15531 | 376 |
NONE => (m $ fst (subst t) $ fst (subst u), NONE) |
377 |
| SOME P => (HOLogic.mk_binop inductive_conj_name (s, P $ t), SOME (s, P $ t))) |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
378 |
| subst s = |
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
379 |
(case pred_of s of |
15531 | 380 |
SOME P => (HOLogic.mk_binop "op Int" |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
381 |
(s, HOLogic.Collect_const (HOLogic.dest_setT |
15531 | 382 |
(fastype_of s)) $ P), NONE) |
383 |
| NONE => (case s of |
|
384 |
(t $ u) => (fst (subst t) $ fst (subst u), NONE) |
|
385 |
| (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE) |
|
386 |
| _ => (s, NONE))); |
|
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
387 |
|
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
388 |
fun mk_prem (s, prems) = (case subst s of |
15531 | 389 |
(_, SOME (t, u)) => t :: u :: prems |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
390 |
| (t, _) => t :: prems); |
9598 | 391 |
|
5094 | 392 |
val Const ("op :", _) $ t $ u = |
393 |
HOLogic.dest_Trueprop (Logic.strip_imp_concl r) |
|
394 |
||
395 |
in list_all_free (frees, |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
396 |
Logic.list_implies (map HOLogic.mk_Trueprop (foldr mk_prem |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
397 |
[] (map HOLogic.dest_Trueprop (Logic.strip_imp_prems r))), |
15570 | 398 |
HOLogic.mk_Trueprop (valOf (pred_of u) $ t))) |
5094 | 399 |
end; |
400 |
||
401 |
val ind_prems = map mk_ind_prem intr_ts; |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
402 |
|
17485 | 403 |
val factors = Library.fold (mg_prod_factors preds) ind_prems []; |
5094 | 404 |
|
405 |
(* make conclusions for induction rules *) |
|
406 |
||
407 |
fun mk_ind_concl ((c, P), (ts, x)) = |
|
408 |
let val T = HOLogic.dest_setT (fastype_of c); |
|
17485 | 409 |
val ps = AList.lookup (op =) factors P |> the_default []; |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
410 |
val Ts = prodT_factors [] ps T; |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
411 |
val (frees, x') = foldr (fn (T', (fs, s)) => |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
412 |
((Free (s, T'))::fs, Symbol.bump_string s)) ([], x) Ts; |
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
413 |
val tuple = mk_tuple [] ps T frees; |
5094 | 414 |
in ((HOLogic.mk_binop "op -->" |
415 |
(HOLogic.mk_mem (tuple, c), P $ tuple))::ts, x') |
|
416 |
end; |
|
417 |
||
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
418 |
val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
419 |
(fst (foldr mk_ind_concl ([], "xa") (cs ~~ preds)))) |
5094 | 420 |
|
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
421 |
in (preds, ind_prems, mutual_ind_concl, |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
422 |
map (apfst (fst o dest_Free)) factors) |
5094 | 423 |
end; |
424 |
||
6424 | 425 |
|
10735 | 426 |
(* prepare cases and induct rules *) |
8316
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
427 |
|
18222 | 428 |
fun add_cases_induct no_elim no_induct coind names elims induct = |
8316
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
429 |
let |
18330 | 430 |
fun cases_spec name elim thy = |
9405 | 431 |
thy |
18463 | 432 |
|> Theory.parent_path |
9405 | 433 |
|> Theory.add_path (Sign.base_name name) |
18728 | 434 |
|> PureThy.add_thms [(("cases", elim), [InductAttrib.cases_set name])] |> snd |
18463 | 435 |
|> Theory.restore_naming thy; |
18330 | 436 |
val cases_specs = if no_elim then [] else map2 cases_spec names elims; |
8316
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
437 |
|
18728 | 438 |
val induct_att = if coind then InductAttrib.coinduct_set else InductAttrib.induct_set; |
19874 | 439 |
fun induct_specs thy = |
440 |
if no_induct then thy |
|
18463 | 441 |
else |
442 |
let |
|
19874 | 443 |
val ctxt = ProofContext.init thy; |
444 |
val rules = names ~~ ProjectRule.projects ctxt (1 upto length names) induct; |
|
18463 | 445 |
val inducts = map (RuleCases.save induct o standard o #2) rules; |
446 |
in |
|
19874 | 447 |
thy |
448 |
|> PureThy.add_thms (rules |> map (fn (name, th) => |
|
449 |
(("", th), [RuleCases.consumes 1, induct_att name]))) |> snd |
|
450 |
|> PureThy.add_thmss |
|
451 |
[((coind_prefix coind ^ "inducts", inducts), [RuleCases.consumes 1])] |> snd |
|
18463 | 452 |
end; |
453 |
in Library.apply cases_specs #> induct_specs end; |
|
8316
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
454 |
|
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
455 |
|
74639e19eca0
add_cases_induct: project_rules accomodates mutual induction;
wenzelm
parents:
8312
diff
changeset
|
456 |
|
10735 | 457 |
(** proofs for (co)inductive sets **) |
6424 | 458 |
|
10735 | 459 |
(* prove monotonicity -- NOT subject to quick_and_dirty! *) |
5094 | 460 |
|
461 |
fun prove_mono setT fp_fun monos thy = |
|
10735 | 462 |
(message " Proving monotonicity ..."; |
20047 | 463 |
Goal.prove_global thy [] [] (*NO quick_and_dirty here!*) |
17985 | 464 |
(HOLogic.mk_Trueprop |
465 |
(Const (mono_name, (setT --> setT) --> HOLogic.boolT) $ fp_fun)) |
|
466 |
(fn _ => EVERY [rtac monoI 1, |
|
20047 | 467 |
REPEAT (ares_tac (List.concat (map mk_mono monos) @ get_monos thy) 1)])); |
5094 | 468 |
|
6424 | 469 |
|
10735 | 470 |
(* prove introduction rules *) |
5094 | 471 |
|
20047 | 472 |
fun prove_intrs coind mono fp_def intr_ts rec_sets_defs ctxt = |
5094 | 473 |
let |
10735 | 474 |
val _ = clean_message " Proving the introduction rules ..."; |
5094 | 475 |
|
13657
c1637d60a846
Now applies standard' to "unfold" theorem (due to flex-flex constraints).
berghofe
parents:
13626
diff
changeset
|
476 |
val unfold = standard' (mono RS (fp_def RS |
10186 | 477 |
(if coind then def_gfp_unfold else def_lfp_unfold))); |
5094 | 478 |
|
479 |
fun select_disj 1 1 = [] |
|
480 |
| select_disj _ 1 = [rtac disjI1] |
|
481 |
| select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1)); |
|
482 |
||
17985 | 483 |
val intrs = (1 upto (length intr_ts) ~~ intr_ts) |> map (fn (i, intr) => |
20047 | 484 |
rulify (SkipProof.prove ctxt [] [] intr (fn _ => EVERY |
17985 | 485 |
[rewrite_goals_tac rec_sets_defs, |
486 |
stac unfold 1, |
|
487 |
REPEAT (resolve_tac [vimageI2, CollectI] 1), |
|
488 |
(*Now 1-2 subgoals: the disjunction, perhaps equality.*) |
|
489 |
EVERY1 (select_disj (length intr_ts) i), |
|
490 |
(*Not ares_tac, since refl must be tried before any equality assumptions; |
|
491 |
backtracking may occur if the premises have extra variables!*) |
|
492 |
DEPTH_SOLVE_1 (resolve_tac [refl, exI, conjI] 1 APPEND assume_tac 1), |
|
493 |
(*Now solve the equations like Inl 0 = Inl ?b2*) |
|
18222 | 494 |
REPEAT (rtac refl 1)]))) |
5094 | 495 |
|
496 |
in (intrs, unfold) end; |
|
497 |
||
6424 | 498 |
|
10735 | 499 |
(* prove elimination rules *) |
5094 | 500 |
|
20047 | 501 |
fun prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs ctxt = |
5094 | 502 |
let |
10735 | 503 |
val _ = clean_message " Proving the elimination rules ..."; |
5094 | 504 |
|
18787
5784fe1b5657
Inductive sets with no introduction rules are now allowed as well.
berghofe
parents:
18728
diff
changeset
|
505 |
val rules1 = [CollectE, disjE, make_elim vimageD, exE, FalseE]; |
10735 | 506 |
val rules2 = [conjE, Inl_neq_Inr, Inr_neq_Inl] @ map make_elim [Inl_inject, Inr_inject]; |
8375 | 507 |
in |
11005 | 508 |
mk_elims cs cTs params intr_ts intr_names |> map (fn (t, cases) => |
20047 | 509 |
SkipProof.prove ctxt [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) |
17985 | 510 |
(fn prems => EVERY |
11005 | 511 |
[cut_facts_tac [hd prems] 1, |
17985 | 512 |
rewrite_goals_tac rec_sets_defs, |
11005 | 513 |
dtac (unfold RS subst) 1, |
514 |
REPEAT (FIRSTGOAL (eresolve_tac rules1)), |
|
515 |
REPEAT (FIRSTGOAL (eresolve_tac rules2)), |
|
17985 | 516 |
EVERY (map (fn prem => |
517 |
DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_sets_defs prem, conjI] 1)) (tl prems))]) |
|
11005 | 518 |
|> rulify |
519 |
|> RuleCases.name cases) |
|
8375 | 520 |
end; |
5094 | 521 |
|
6424 | 522 |
|
10735 | 523 |
(* derivation of simplified elimination rules *) |
5094 | 524 |
|
11682
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
525 |
local |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
526 |
|
7107 | 527 |
(*cprop should have the form t:Si where Si is an inductive set*) |
11682
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
528 |
val mk_cases_err = "mk_cases: proposition not of form \"t : S_i\""; |
9598 | 529 |
|
11682
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
530 |
(*delete needless equality assumptions*) |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
531 |
val refl_thin = prove_goal HOL.thy "!!P. a = a ==> P ==> P" (fn _ => [assume_tac 1]); |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
532 |
val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE, Pair_inject]; |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
533 |
val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls; |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
534 |
|
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
535 |
fun simp_case_tac solved ss i = |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
536 |
EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
537 |
THEN_MAYBE (if solved then no_tac else all_tac); |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
538 |
|
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
539 |
in |
9598 | 540 |
|
541 |
fun mk_cases_i elims ss cprop = |
|
7107 | 542 |
let |
543 |
val prem = Thm.assume cprop; |
|
11682
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
544 |
val tac = ALLGOALS (simp_case_tac false ss) THEN prune_params_tac; |
9298 | 545 |
fun mk_elim rl = Drule.standard (Tactic.rule_by_tactic tac (prem RS rl)); |
7107 | 546 |
in |
547 |
(case get_first (try mk_elim) elims of |
|
15531 | 548 |
SOME r => r |
549 |
| NONE => error (Pretty.string_of (Pretty.block |
|
9598 | 550 |
[Pretty.str mk_cases_err, Pretty.fbrk, Display.pretty_cterm cprop]))) |
7107 | 551 |
end; |
552 |
||
6141 | 553 |
fun mk_cases elims s = |
16432 | 554 |
mk_cases_i elims (simpset()) (Thm.read_cterm (Thm.theory_of_thm (hd elims)) (s, propT)); |
9598 | 555 |
|
556 |
fun smart_mk_cases thy ss cprop = |
|
557 |
let |
|
558 |
val c = #1 (Term.dest_Const (Term.head_of (#2 (HOLogic.dest_mem (HOLogic.dest_Trueprop |
|
559 |
(Logic.strip_imp_concl (Thm.term_of cprop))))))) handle TERM _ => error mk_cases_err; |
|
560 |
val (_, {elims, ...}) = the_inductive thy c; |
|
561 |
in mk_cases_i elims ss cprop end; |
|
7107 | 562 |
|
11682
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
563 |
end; |
d9063229b4a1
simp_case_tac is back again from induct_method.ML;
wenzelm
parents:
11628
diff
changeset
|
564 |
|
7107 | 565 |
|
566 |
(* inductive_cases(_i) *) |
|
567 |
||
12609 | 568 |
fun gen_inductive_cases prep_att prep_prop args thy = |
9598 | 569 |
let |
16432 | 570 |
val cert_prop = Thm.cterm_of thy o prep_prop (ProofContext.init thy); |
12609 | 571 |
val mk_cases = smart_mk_cases thy (Simplifier.simpset_of thy) o cert_prop; |
572 |
||
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12798
diff
changeset
|
573 |
val facts = args |> map (fn ((a, atts), props) => |
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12798
diff
changeset
|
574 |
((a, map (prep_att thy) atts), map (Thm.no_attributes o single o mk_cases) props)); |
18799 | 575 |
in thy |> IsarThy.theorems_i PureThy.lemmaK facts |> snd end; |
5094 | 576 |
|
18728 | 577 |
val inductive_cases = gen_inductive_cases Attrib.attribute ProofContext.read_prop; |
12172 | 578 |
val inductive_cases_i = gen_inductive_cases (K I) ProofContext.cert_prop; |
7107 | 579 |
|
6424 | 580 |
|
9598 | 581 |
(* mk_cases_meth *) |
582 |
||
583 |
fun mk_cases_meth (ctxt, raw_props) = |
|
584 |
let |
|
585 |
val thy = ProofContext.theory_of ctxt; |
|
15032 | 586 |
val ss = local_simpset_of ctxt; |
16432 | 587 |
val cprops = map (Thm.cterm_of thy o ProofContext.read_prop ctxt) raw_props; |
10743 | 588 |
in Method.erule 0 (map (smart_mk_cases thy ss) cprops) end; |
9598 | 589 |
|
590 |
val mk_cases_args = Method.syntax (Scan.lift (Scan.repeat1 Args.name)); |
|
591 |
||
592 |
||
10735 | 593 |
(* prove induction rule *) |
5094 | 594 |
|
595 |
fun prove_indrule cs cTs sumT rec_const params intr_ts mono |
|
20047 | 596 |
fp_def rec_sets_defs ctxt = |
5094 | 597 |
let |
10735 | 598 |
val _ = clean_message " Proving the induction rule ..."; |
20047 | 599 |
val thy = ProofContext.theory_of ctxt; |
5094 | 600 |
|
12922 | 601 |
val sum_case_rewrites = |
16432 | 602 |
(if Context.theory_name thy = "Datatype" then |
16486 | 603 |
PureThy.get_thms thy (Name "sum.cases") |
16432 | 604 |
else |
605 |
(case ThyInfo.lookup_theory "Datatype" of |
|
606 |
NONE => [] |
|
16975 | 607 |
| SOME thy' => |
608 |
if Theory.subthy (thy', thy) then |
|
609 |
PureThy.get_thms thy' (Name "sum.cases") |
|
610 |
else [])) |
|
16432 | 611 |
|> map mk_meta_eq; |
7293 | 612 |
|
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
613 |
val (preds, ind_prems, mutual_ind_concl, factors) = |
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
614 |
mk_indrule cs cTs params intr_ts; |
5094 | 615 |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
616 |
val dummy = if !trace then |
17985 | 617 |
(writeln "ind_prems = "; |
618 |
List.app (writeln o Sign.string_of_term thy) ind_prems) |
|
619 |
else (); |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
620 |
|
5094 | 621 |
(* make predicate for instantiation of abstract induction rule *) |
622 |
||
623 |
fun mk_ind_pred _ [P] = P |
|
624 |
| mk_ind_pred T Ps = |
|
625 |
let val n = (length Ps) div 2; |
|
626 |
val Type (_, [T1, T2]) = T |
|
7293 | 627 |
in Const ("Datatype.sum.sum_case", |
5094 | 628 |
[T1 --> HOLogic.boolT, T2 --> HOLogic.boolT, T] ---> HOLogic.boolT) $ |
15570 | 629 |
mk_ind_pred T1 (Library.take (n, Ps)) $ mk_ind_pred T2 (Library.drop (n, Ps)) |
5094 | 630 |
end; |
631 |
||
632 |
val ind_pred = mk_ind_pred sumT preds; |
|
633 |
||
634 |
val ind_concl = HOLogic.mk_Trueprop |
|
635 |
(HOLogic.all_const sumT $ Abs ("x", sumT, HOLogic.mk_binop "op -->" |
|
636 |
(HOLogic.mk_mem (Bound 0, rec_const), ind_pred $ Bound 0))); |
|
637 |
||
638 |
(* simplification rules for vimage and Collect *) |
|
639 |
||
640 |
val vimage_simps = if length cs < 2 then [] else |
|
20047 | 641 |
map (fn c => standard (SkipProof.prove ctxt [] [] |
5094 | 642 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq |
643 |
(mk_vimage cs sumT (HOLogic.Collect_const sumT $ ind_pred) c, |
|
644 |
HOLogic.Collect_const (HOLogic.dest_setT (fastype_of c)) $ |
|
17985 | 645 |
List.nth (preds, find_index_eq c cs)))) |
646 |
(fn _ => EVERY |
|
647 |
[rtac vimage_Collect 1, rewrite_goals_tac sum_case_rewrites, rtac refl 1]))) cs; |
|
5094 | 648 |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
649 |
val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct)); |
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
650 |
|
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
651 |
val dummy = if !trace then |
17985 | 652 |
(writeln "raw_fp_induct = "; print_thm raw_fp_induct) |
653 |
else (); |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
654 |
|
20047 | 655 |
val induct = standard (SkipProof.prove ctxt [] ind_prems ind_concl |
17985 | 656 |
(fn prems => EVERY |
657 |
[rewrite_goals_tac [inductive_conj_def], |
|
658 |
rtac (impI RS allI) 1, |
|
13626
282fbabec862
Fixed bug involving inductive definitions having equalities in the premises,
paulson
parents:
13197
diff
changeset
|
659 |
DETERM (etac raw_fp_induct 1), |
7710
bf8cb3fc5d64
Monotonicity rules for inductive definitions can now be added to a theory via
berghofe
parents:
7349
diff
changeset
|
660 |
rewrite_goals_tac (map mk_meta_eq (vimage_Int::Int_Collect::vimage_simps)), |
5094 | 661 |
fold_goals_tac rec_sets_defs, |
662 |
(*This CollectE and disjE separates out the introduction rules*) |
|
18787
5784fe1b5657
Inductive sets with no introduction rules are now allowed as well.
berghofe
parents:
18728
diff
changeset
|
663 |
REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE, exE, FalseE])), |
5094 | 664 |
(*Now break down the individual cases. No disjE here in case |
665 |
some premise involves disjunction.*) |
|
13747
bf308fcfd08e
Better treatment of equality in premises of inductive definitions. Less
paulson
parents:
13709
diff
changeset
|
666 |
REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)), |
19359
5d523a1b6ddc
Fixed bug that caused proof of induction rule to fail
berghofe
parents:
18921
diff
changeset
|
667 |
rewrite_goals_tac sum_case_rewrites, |
5094 | 668 |
EVERY (map (fn prem => |
17985 | 669 |
DEPTH_SOLVE_1 (ares_tac [rewrite_rule [inductive_conj_def] prem, conjI, refl] 1)) prems)])); |
5094 | 670 |
|
20047 | 671 |
val lemma = standard (SkipProof.prove ctxt [] [] |
17985 | 672 |
(Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY |
673 |
[rewrite_goals_tac rec_sets_defs, |
|
5094 | 674 |
REPEAT (EVERY |
675 |
[REPEAT (resolve_tac [conjI, impI] 1), |
|
676 |
TRY (dtac vimageD 1), etac allE 1, dtac mp 1, atac 1, |
|
7293 | 677 |
rewrite_goals_tac sum_case_rewrites, |
17985 | 678 |
atac 1])])) |
5094 | 679 |
|
10988
e0016a009c17
Splitting of arguments of product types in induction rules is now less
berghofe
parents:
10910
diff
changeset
|
680 |
in standard (split_rule factors (induct RS lemma)) end; |
5094 | 681 |
|
6424 | 682 |
|
683 |
||
10735 | 684 |
(** specification of (co)inductive sets **) |
5094 | 685 |
|
10729 | 686 |
fun cond_declare_consts declare_consts cs paramTs cnames = |
687 |
if declare_consts then |
|
14235
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
688 |
Theory.add_consts_i (map (fn (c, n) => (Sign.base_name n, paramTs ---> fastype_of c, NoSyn)) (cs ~~ cnames)) |
10729 | 689 |
else I; |
690 |
||
12180 | 691 |
fun mk_ind_def declare_consts alt_name coind cs intr_ts monos thy |
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
692 |
params paramTs cTs cnames = |
5094 | 693 |
let |
694 |
val sumT = fold_bal (fn (T, U) => Type ("+", [T, U])) cTs; |
|
695 |
val setT = HOLogic.mk_setT sumT; |
|
696 |
||
10735 | 697 |
val fp_name = if coind then gfp_name else lfp_name; |
5094 | 698 |
|
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
699 |
val used = foldr add_term_names [] intr_ts; |
20071
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
20047
diff
changeset
|
700 |
val [sname, xname] = Name.variant_list used ["S", "x"]; |
5149 | 701 |
|
5094 | 702 |
(* transform an introduction rule into a conjunction *) |
703 |
(* [| t : ... S_i ... ; ... |] ==> u : S_j *) |
|
704 |
(* is transformed into *) |
|
705 |
(* x = Inj_j u & t : ... Inj_i -`` S ... & ... *) |
|
706 |
||
707 |
fun transform_rule r = |
|
708 |
let |
|
709 |
val frees = map dest_Free ((add_term_frees (r, [])) \\ params); |
|
5149 | 710 |
val subst = subst_free |
711 |
(cs ~~ (map (mk_vimage cs sumT (Free (sname, setT))) cs)); |
|
5094 | 712 |
val Const ("op :", _) $ t $ u = |
713 |
HOLogic.dest_Trueprop (Logic.strip_imp_concl r) |
|
714 |
||
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
715 |
in foldr (fn ((x, T), P) => HOLogic.mk_exists (x, T, P)) |
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
716 |
(foldr1 HOLogic.mk_conj |
5149 | 717 |
(((HOLogic.eq_const sumT) $ Free (xname, sumT) $ (mk_inj cs sumT u t)):: |
5094 | 718 |
(map (subst o HOLogic.dest_Trueprop) |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
719 |
(Logic.strip_imp_prems r)))) frees |
5094 | 720 |
end |
721 |
||
722 |
(* make a disjunction of all introduction rules *) |
|
723 |
||
5149 | 724 |
val fp_fun = absfree (sname, setT, (HOLogic.Collect_const sumT) $ |
18787
5784fe1b5657
Inductive sets with no introduction rules are now allowed as well.
berghofe
parents:
18728
diff
changeset
|
725 |
absfree (xname, sumT, if null intr_ts then HOLogic.false_const |
5784fe1b5657
Inductive sets with no introduction rules are now allowed as well.
berghofe
parents:
18728
diff
changeset
|
726 |
else foldr1 HOLogic.mk_disj (map transform_rule intr_ts))); |
5094 | 727 |
|
728 |
(* add definiton of recursive sets to theory *) |
|
729 |
||
14235
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
730 |
val rec_name = if alt_name = "" then |
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
731 |
space_implode "_" (map Sign.base_name cnames) else alt_name; |
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
732 |
val full_rec_name = if length cs < 2 then hd cnames |
16432 | 733 |
else Sign.full_name thy rec_name; |
5094 | 734 |
|
735 |
val rec_const = list_comb |
|
736 |
(Const (full_rec_name, paramTs ---> setT), params); |
|
737 |
||
738 |
val fp_def_term = Logic.mk_equals (rec_const, |
|
10735 | 739 |
Const (fp_name, (setT --> setT) --> setT) $ fp_fun); |
5094 | 740 |
|
741 |
val def_terms = fp_def_term :: (if length cs < 2 then [] else |
|
742 |
map (fn c => Logic.mk_equals (c, mk_vimage cs sumT rec_const c)) cs); |
|
743 |
||
18358 | 744 |
val ([fp_def :: rec_sets_defs], thy') = |
8433 | 745 |
thy |
10729 | 746 |
|> cond_declare_consts declare_consts cs paramTs cnames |
8433 | 747 |
|> (if length cs < 2 then I |
748 |
else Theory.add_consts_i [(rec_name, paramTs ---> setT, NoSyn)]) |
|
749 |
|> Theory.add_path rec_name |
|
9315 | 750 |
|> PureThy.add_defss_i false [(("defs", def_terms), [])]; |
5094 | 751 |
|
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
752 |
val mono = prove_mono setT fp_fun monos thy' |
5094 | 753 |
|
18222 | 754 |
in (thy', rec_name, mono, fp_def, rec_sets_defs, rec_const, sumT) end; |
5094 | 755 |
|
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
756 |
fun add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs |
12180 | 757 |
intros monos thy params paramTs cTs cnames induct_cases = |
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
758 |
let |
10735 | 759 |
val _ = |
760 |
if verbose then message ("Proofs for " ^ coind_prefix coind ^ "inductive set(s) " ^ |
|
14235
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
761 |
commas_quote (map Sign.base_name cnames)) else (); |
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
762 |
|
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
763 |
val ((intr_names, intr_ts), intr_atts) = apfst split_list (split_list intros); |
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
764 |
|
18222 | 765 |
val (thy1, rec_name, mono, fp_def, rec_sets_defs, rec_const, sumT) = |
12180 | 766 |
mk_ind_def declare_consts alt_name coind cs intr_ts monos thy |
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
767 |
params paramTs cTs cnames; |
20047 | 768 |
val ctxt1 = ProofContext.init thy1; |
9072
a4896cf23638
Now also proves monotonicity when in quick_and_dirty mode.
berghofe
parents:
8720
diff
changeset
|
769 |
|
20047 | 770 |
val (intrs, unfold) = prove_intrs coind mono fp_def intr_ts rec_sets_defs ctxt1; |
5094 | 771 |
val elims = if no_elim then [] else |
20047 | 772 |
prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs ctxt1; |
8312
b470bc28b59d
add_cases_induct: accomodate no_elim and no_ind flags;
wenzelm
parents:
8307
diff
changeset
|
773 |
val raw_induct = if no_ind then Drule.asm_rl else |
5094 | 774 |
if coind then standard (rule_by_tactic |
5553 | 775 |
(rewrite_tac [mk_meta_eq vimage_Un] THEN |
5094 | 776 |
fold_tac rec_sets_defs) (mono RS (fp_def RS def_Collect_coinduct))) |
777 |
else |
|
778 |
prove_indrule cs cTs sumT rec_const params intr_ts mono fp_def |
|
20047 | 779 |
rec_sets_defs ctxt1; |
12165 | 780 |
val induct = |
18222 | 781 |
if coind then |
782 |
(raw_induct, [RuleCases.case_names [rec_name], |
|
18234 | 783 |
RuleCases.case_conclusion (rec_name, induct_cases), |
18222 | 784 |
RuleCases.consumes 1]) |
785 |
else if no_ind orelse length cs > 1 then |
|
786 |
(raw_induct, [RuleCases.case_names induct_cases, RuleCases.consumes 0]) |
|
787 |
else (raw_induct RSN (2, rev_mp), [RuleCases.case_names induct_cases, RuleCases.consumes 1]); |
|
5094 | 788 |
|
18377 | 789 |
val (intrs', thy2) = |
790 |
thy1 |
|
791 |
|> PureThy.add_thms ((intr_names ~~ intrs) ~~ intr_atts); |
|
792 |
val (([_, elims'], [induct']), thy3) = |
|
10735 | 793 |
thy2 |
11005 | 794 |
|> PureThy.add_thmss |
11628 | 795 |
[(("intros", intrs'), []), |
11005 | 796 |
(("elims", elims), [RuleCases.consumes 1])] |
18377 | 797 |
||>> PureThy.add_thms |
18463 | 798 |
[((coind_prefix coind ^ "induct", rulify (#1 induct)), #2 induct)]; |
9939 | 799 |
in (thy3, |
10735 | 800 |
{defs = fp_def :: rec_sets_defs, |
5094 | 801 |
mono = mono, |
802 |
unfold = unfold, |
|
13709 | 803 |
intrs = intrs', |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7710
diff
changeset
|
804 |
elims = elims', |
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7710
diff
changeset
|
805 |
mk_cases = mk_cases elims', |
10729 | 806 |
raw_induct = rulify raw_induct, |
7798
42e94b618f34
return stored thms with proper naming in derivation;
wenzelm
parents:
7710
diff
changeset
|
807 |
induct = induct'}) |
5094 | 808 |
end; |
809 |
||
6424 | 810 |
|
10735 | 811 |
(* external interfaces *) |
5094 | 812 |
|
16432 | 813 |
fun try_term f msg thy t = |
10735 | 814 |
(case Library.try f t of |
15531 | 815 |
SOME x => x |
16432 | 816 |
| NONE => error (msg ^ Sign.string_of_term thy t)); |
5094 | 817 |
|
12180 | 818 |
fun add_inductive_i verbose declare_consts alt_name coind no_elim no_ind cs pre_intros monos thy = |
5094 | 819 |
let |
6424 | 820 |
val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions"); |
5094 | 821 |
|
822 |
(*parameters should agree for all mutually recursive components*) |
|
823 |
val (_, params) = strip_comb (hd cs); |
|
10735 | 824 |
val paramTs = map (try_term (snd o dest_Free) "Parameter in recursive\ |
16432 | 825 |
\ component is not a free variable: " thy) params; |
5094 | 826 |
|
10735 | 827 |
val cTs = map (try_term (HOLogic.dest_setT o fastype_of) |
16432 | 828 |
"Recursive component not of type set: " thy) cs; |
5094 | 829 |
|
14235
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
830 |
val cnames = map (try_term (fst o dest_Const o head_of) |
16432 | 831 |
"Recursive set not previously declared as constant: " thy) cs; |
5094 | 832 |
|
16432 | 833 |
val save_thy = thy |
834 |
|> Theory.copy |> cond_declare_consts declare_consts cs paramTs cnames; |
|
835 |
val intros = map (check_rule save_thy cs) pre_intros; |
|
8401 | 836 |
val induct_cases = map (#1 o #1) intros; |
6437 | 837 |
|
9405 | 838 |
val (thy1, result as {elims, induct, ...}) = |
11628 | 839 |
add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs intros monos |
12180 | 840 |
thy params paramTs cTs cnames induct_cases; |
8307 | 841 |
val thy2 = thy1 |
14235
281295a1bbaa
Fixed bug in mk_ind_def that caused the inductive definition package to
berghofe
parents:
13747
diff
changeset
|
842 |
|> put_inductives cnames ({names = cnames, coind = coind}, result) |
18463 | 843 |
|> add_cases_induct no_elim no_ind coind cnames elims induct |
844 |
|> Theory.parent_path; |
|
6437 | 845 |
in (thy2, result) end; |
5094 | 846 |
|
12180 | 847 |
fun add_inductive verbose coind c_strings intro_srcs raw_monos thy = |
5094 | 848 |
let |
16975 | 849 |
val cs = map (Sign.read_term thy) c_strings; |
6424 | 850 |
|
851 |
val intr_names = map (fst o fst) intro_srcs; |
|
16432 | 852 |
fun read_rule s = Thm.read_cterm thy (s, propT) |
18678 | 853 |
handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s); |
9405 | 854 |
val intr_ts = map (Thm.term_of o read_rule o snd o fst) intro_srcs; |
18728 | 855 |
val intr_atts = map (map (Attrib.attribute thy) o snd) intro_srcs; |
16432 | 856 |
val (cs', intr_ts') = unify_consts thy cs intr_ts; |
5094 | 857 |
|
18418
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
858 |
val (monos, thy') = thy |> IsarThy.apply_theorems raw_monos; |
6424 | 859 |
in |
7020
75ff179df7b7
Exported function unify_consts (workaround to avoid inconsistently
berghofe
parents:
6851
diff
changeset
|
860 |
add_inductive_i verbose false "" coind false false cs' |
12180 | 861 |
((intr_names ~~ intr_ts') ~~ intr_atts) monos thy' |
5094 | 862 |
end; |
863 |
||
6424 | 864 |
|
865 |
||
6437 | 866 |
(** package setup **) |
867 |
||
868 |
(* setup theory *) |
|
869 |
||
8634 | 870 |
val setup = |
18708 | 871 |
InductiveData.init #> |
9625 | 872 |
Method.add_methods [("ind_cases", mk_cases_meth oo mk_cases_args, |
18708 | 873 |
"dynamic case analysis on sets")] #> |
18728 | 874 |
Attrib.add_attributes [("mono", Attrib.add_del_args mono_add mono_del, |
875 |
"declaration of monotonicity rule")]; |
|
6437 | 876 |
|
877 |
||
878 |
(* outer syntax *) |
|
6424 | 879 |
|
17057 | 880 |
local structure P = OuterParse and K = OuterKeyword in |
6424 | 881 |
|
12180 | 882 |
fun mk_ind coind ((sets, intrs), monos) = |
883 |
#1 o add_inductive true coind sets (map P.triple_swap intrs) monos; |
|
6424 | 884 |
|
885 |
fun ind_decl coind = |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12798
diff
changeset
|
886 |
Scan.repeat1 P.term -- |
9598 | 887 |
(P.$$$ "intros" |-- |
18787
5784fe1b5657
Inductive sets with no introduction rules are now allowed as well.
berghofe
parents:
18728
diff
changeset
|
888 |
P.!!! (Scan.repeat (P.opt_thm_name ":" -- P.prop))) -- |
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12798
diff
changeset
|
889 |
Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) [] |
6424 | 890 |
>> (Toplevel.theory o mk_ind coind); |
891 |
||
6723 | 892 |
val inductiveP = |
893 |
OuterSyntax.command "inductive" "define inductive sets" K.thy_decl (ind_decl false); |
|
894 |
||
895 |
val coinductiveP = |
|
896 |
OuterSyntax.command "coinductive" "define coinductive sets" K.thy_decl (ind_decl true); |
|
6424 | 897 |
|
7107 | 898 |
|
899 |
val ind_cases = |
|
12876
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents:
12798
diff
changeset
|
900 |
P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 P.prop) |
7107 | 901 |
>> (Toplevel.theory o inductive_cases); |
902 |
||
903 |
val inductive_casesP = |
|
9804 | 904 |
OuterSyntax.command "inductive_cases" |
9598 | 905 |
"create simplified instances of elimination rules (improper)" K.thy_script ind_cases; |
7107 | 906 |
|
12180 | 907 |
val _ = OuterSyntax.add_keywords ["intros", "monos"]; |
7107 | 908 |
val _ = OuterSyntax.add_parsers [inductiveP, coinductiveP, inductive_casesP]; |
6424 | 909 |
|
5094 | 910 |
end; |
6424 | 911 |
|
912 |
end; |
|
15705 | 913 |