author | paulson |
Fri, 19 Sep 1997 16:12:21 +0200 | |
changeset 3685 | 5b8c0c8f576e |
parent 2688 | 889a1cbd1aca |
child 3978 | 7e1cfed19d94 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/intr_elim.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Introduction/elimination rule module -- for Inductive/Coinductive Definitions |
|
7 |
*) |
|
8 |
||
1465 | 9 |
signature INDUCTIVE_ARG = (** Description of a (co)inductive def **) |
923 | 10 |
sig |
11 |
val thy : theory (*new theory with inductive defs*) |
|
1465 | 12 |
val monos : thm list (*monotonicity of each M operator*) |
13 |
val con_defs : thm list (*definitions of the constructors*) |
|
923 | 14 |
end; |
15 |
||
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
16 |
|
1465 | 17 |
signature INDUCTIVE_I = (** Terms read from the theory section **) |
923 | 18 |
sig |
1465 | 19 |
val rec_tms : term list (*the recursive sets*) |
20 |
val intr_tms : term list (*terms for the introduction rules*) |
|
923 | 21 |
end; |
22 |
||
23 |
signature INTR_ELIM = |
|
24 |
sig |
|
25 |
val thy : theory (*copy of input theory*) |
|
1465 | 26 |
val defs : thm list (*definitions made in thy*) |
27 |
val mono : thm (*monotonicity for the lfp definition*) |
|
28 |
val intrs : thm list (*introduction rules*) |
|
29 |
val elim : thm (*case analysis theorem*) |
|
30 |
val mk_cases : thm list -> string -> thm (*generates case theorems*) |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
31 |
end; |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
32 |
|
1465 | 33 |
signature INTR_ELIM_AUX = (** Used to make induction rules **) |
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
34 |
sig |
1465 | 35 |
val raw_induct : thm (*raw induction rule from Fp.induct*) |
36 |
val rec_names : string list (*names of recursive sets*) |
|
923 | 37 |
end; |
38 |
||
39 |
(*prove intr/elim rules for a fixedpoint definition*) |
|
40 |
functor Intr_elim_Fun |
|
41 |
(structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
42 |
and Fp: FP) |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
43 |
: sig include INTR_ELIM INTR_ELIM_AUX end = |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
44 |
let |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
45 |
val rec_names = map (#1 o dest_Const o head_of) Inductive.rec_tms; |
923 | 46 |
val big_rec_name = space_implode "_" rec_names; |
47 |
||
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
48 |
val _ = deny (big_rec_name mem map ! (stamps_of_thy Inductive.thy)) |
923 | 49 |
("Definition " ^ big_rec_name ^ |
1465 | 50 |
" would clash with the theory of the same name!"); |
923 | 51 |
|
52 |
(*fetch fp definitions from the theory*) |
|
53 |
val big_rec_def::part_rec_defs = |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
54 |
map (get_def Inductive.thy) |
923 | 55 |
(case rec_names of [_] => rec_names | _ => big_rec_name::rec_names); |
56 |
||
57 |
||
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
58 |
val sign = sign_of Inductive.thy; |
923 | 59 |
|
60 |
(********) |
|
61 |
val _ = writeln " Proving monotonicity..."; |
|
62 |
||
63 |
val Const("==",_) $ _ $ (Const(_,fpT) $ fp_abs) = |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
64 |
big_rec_def |> rep_thm |> #prop |> Logic.unvarify; |
923 | 65 |
|
66 |
(*For the type of the argument of mono*) |
|
67 |
val [monoT] = binder_types fpT; |
|
68 |
||
69 |
val mono = |
|
70 |
prove_goalw_cterm [] |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
71 |
(cterm_of sign (Ind_Syntax.mk_Trueprop |
1465 | 72 |
(Const("mono", monoT --> Ind_Syntax.boolT) $ fp_abs))) |
923 | 73 |
(fn _ => |
74 |
[rtac monoI 1, |
|
1465 | 75 |
REPEAT (ares_tac (basic_monos @ Inductive.monos) 1)]); |
923 | 76 |
|
77 |
val unfold = standard (mono RS (big_rec_def RS Fp.Tarski)); |
|
78 |
||
79 |
(********) |
|
80 |
val _ = writeln " Proving the introduction rules..."; |
|
81 |
||
82 |
fun intro_tacsf disjIn prems = |
|
83 |
[(*insert prems and underlying sets*) |
|
84 |
cut_facts_tac prems 1, |
|
2031 | 85 |
stac unfold 1, |
923 | 86 |
REPEAT (resolve_tac [Part_eqI,CollectI] 1), |
87 |
(*Now 1-2 subgoals: the disjunction, perhaps equality.*) |
|
88 |
rtac disjIn 1, |
|
89 |
(*Not ares_tac, since refl must be tried before any equality assumptions; |
|
90 |
backtracking may occur if the premises have extra variables!*) |
|
2414
13df7d6c5c3b
intro_tacsf: replaced ORELSE by APPEND in order to stop
paulson
parents:
2270
diff
changeset
|
91 |
DEPTH_SOLVE_1 (resolve_tac [refl,exI,conjI] 1 APPEND assume_tac 1), |
1191
d7e0c280f261
Added two final lines to intro_tacsf for mutual recursion
lcp
parents:
923
diff
changeset
|
92 |
(*Now solve the equations like Inl 0 = Inl ?b2*) |
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
93 |
rewrite_goals_tac Inductive.con_defs, |
1191
d7e0c280f261
Added two final lines to intro_tacsf for mutual recursion
lcp
parents:
923
diff
changeset
|
94 |
REPEAT (rtac refl 1)]; |
d7e0c280f261
Added two final lines to intro_tacsf for mutual recursion
lcp
parents:
923
diff
changeset
|
95 |
|
923 | 96 |
|
97 |
(*combines disjI1 and disjI2 to access the corresponding nested disjunct...*) |
|
98 |
val mk_disj_rls = |
|
99 |
let fun f rl = rl RS disjI1 |
|
1465 | 100 |
and g rl = rl RS disjI2 |
923 | 101 |
in accesses_bal(f, g, asm_rl) end; |
102 |
||
2270 | 103 |
val intrs = ListPair.map (uncurry (prove_goalw_cterm part_rec_defs)) |
104 |
(map (cterm_of sign) Inductive.intr_tms, |
|
1465 | 105 |
map intro_tacsf (mk_disj_rls(length Inductive.intr_tms))); |
923 | 106 |
|
107 |
(********) |
|
108 |
val _ = writeln " Proving the elimination rule..."; |
|
109 |
||
110 |
(*Breaks down logical connectives in the monotonic function*) |
|
111 |
val basic_elim_tac = |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
112 |
REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ |
1465 | 113 |
Ind_Syntax.sumprod_free_SEs) |
114 |
ORELSE' bound_hyp_subst_tac)) |
|
923 | 115 |
THEN prune_params_tac; |
116 |
||
117 |
(*Applies freeness of the given constructors, which *must* be unfolded by |
|
118 |
the given defs. Cannot simply use the local con_defs because con_defs=[] |
|
119 |
for inference systems. |
|
120 |
fun con_elim_tac defs = |
|
121 |
rewrite_goals_tac defs THEN basic_elim_tac THEN fold_tac defs; |
|
122 |
*) |
|
123 |
fun con_elim_tac simps = |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
124 |
let val elim_tac = REPEAT o (eresolve_tac (Ind_Syntax.elim_rls @ |
1465 | 125 |
Ind_Syntax.sumprod_free_SEs)) |
923 | 126 |
in ALLGOALS(EVERY'[elim_tac, |
1264
3eb91524b938
added local simpsets; removed IOA from 'make test'
clasohm
parents:
1191
diff
changeset
|
127 |
asm_full_simp_tac (simpset_of "Nat" addsimps simps), |
923 | 128 |
elim_tac, |
129 |
REPEAT o bound_hyp_subst_tac]) |
|
130 |
THEN prune_params_tac |
|
131 |
end; |
|
132 |
||
133 |
||
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
134 |
in |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
135 |
struct |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
136 |
val thy = Inductive.thy |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
137 |
and defs = big_rec_def :: part_rec_defs |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
138 |
and mono = mono |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
139 |
and intrs = intrs; |
923 | 140 |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
141 |
val elim = rule_by_tactic basic_elim_tac |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
142 |
(unfold RS Ind_Syntax.equals_CollectD); |
923 | 143 |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
144 |
(*String s should have the form t:Si where Si is an inductive set*) |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
145 |
fun mk_cases defs s = |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
146 |
rule_by_tactic (con_elim_tac defs) |
2688 | 147 |
(assume_read Inductive.thy s RS elim) |
148 |
|> standard; |
|
1425
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
149 |
|
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
150 |
val raw_induct = standard ([big_rec_def, mono] MRS Fp.induct) |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
151 |
and rec_names = rec_names |
7b61bcfeaa95
Removed unfold:thm from signature INTR_ELIM and from the
paulson
parents:
1264
diff
changeset
|
152 |
end |
923 | 153 |
end; |
154 |