author | wenzelm |
Wed, 29 Apr 1998 11:46:42 +0200 | |
changeset 4876 | 1c502a82bcde |
parent 4871 | fe076613e122 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/add_ind_def.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Fixedpoint definition module -- for Inductive/Coinductive Definitions |
|
7 |
||
8 |
Features: |
|
9 |
* least or greatest fixedpoints |
|
10 |
* user-specified product and sum constructions |
|
11 |
* mutually recursive definitions |
|
12 |
* definitions involving arbitrary monotone operators |
|
13 |
* automatically proves introduction and elimination rules |
|
14 |
||
15 |
The recursive sets must *already* be declared as constants in parent theory! |
|
16 |
||
17 |
Introduction rules have the form |
|
18 |
[| ti:M(Sj), ..., P(x), ... |] ==> t: Sk |] |
|
19 |
where M is some monotone operator (usually the identity) |
|
20 |
P(x) is any (non-conjunctive) side condition on the free variables |
|
21 |
ti, t are any terms |
|
22 |
Sj, Sk are two of the sets being defined in mutual recursion |
|
23 |
||
24 |
Sums are used only for mutual recursion; |
|
25 |
Products are used only to derive "streamlined" induction rules for relations |
|
26 |
||
27 |
Nestings of disjoint sum types: |
|
28 |
(a+(b+c)) for 3, ((a+b)+(c+d)) for 4, ((a+b)+(c+(d+e))) for 5, |
|
29 |
((a+(b+c))+(d+(e+f))) for 6 |
|
30 |
*) |
|
31 |
||
1465 | 32 |
signature FP = (** Description of a fixed point operator **) |
923 | 33 |
sig |
2859 | 34 |
val checkThy : theory -> unit (*signals error if Lfp/Gfp is missing*) |
1465 | 35 |
val oper : string * typ * term -> term (*fixed point operator*) |
2859 | 36 |
val Tarski : thm (*Tarski's fixed point theorem*) |
37 |
val induct : thm (*induction/coinduction rule*) |
|
923 | 38 |
end; |
39 |
||
40 |
||
41 |
signature ADD_INDUCTIVE_DEF = |
|
42 |
sig |
|
43 |
val add_fp_def_i : term list * term list -> theory -> theory |
|
44 |
end; |
|
45 |
||
46 |
||
47 |
||
48 |
(*Declares functions to add fixedpoint/constructor defs to a theory*) |
|
49 |
functor Add_inductive_def_Fun (Fp: FP) : ADD_INDUCTIVE_DEF = |
|
50 |
struct |
|
1397 | 51 |
open Ind_Syntax; |
923 | 52 |
|
53 |
(*internal version*) |
|
54 |
fun add_fp_def_i (rec_tms, intr_tms) thy = |
|
55 |
let |
|
2859 | 56 |
val dummy = Fp.checkThy thy (*has essential ancestors?*) |
57 |
||
923 | 58 |
val sign = sign_of thy; |
59 |
||
1189 | 60 |
(*rec_params should agree for all mutually recursive components*) |
923 | 61 |
val rec_hds = map head_of rec_tms; |
62 |
||
63 |
val _ = assert_all is_Const rec_hds |
|
1465 | 64 |
(fn t => "Recursive set not previously declared as constant: " ^ |
65 |
Sign.string_of_term sign t); |
|
923 | 66 |
|
67 |
(*Now we know they are all Consts, so get their names, type and params*) |
|
68 |
val rec_names = map (#1 o dest_Const) rec_hds |
|
69 |
and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); |
|
70 |
||
3945 | 71 |
val rec_base_names = map Sign.base_name rec_names; |
72 |
val _ = assert_all Syntax.is_identifier rec_base_names |
|
73 |
(fn a => "Base name of recursive set not an identifier: " ^ a); |
|
923 | 74 |
|
75 |
local (*Checking the introduction rules*) |
|
76 |
val intr_sets = map (#2 o rule_concl_msg sign) intr_tms; |
|
77 |
fun intr_ok set = |
|
1465 | 78 |
case head_of set of Const(a,_) => a mem rec_names | _ => false; |
923 | 79 |
in |
80 |
val _ = assert_all intr_ok intr_sets |
|
1465 | 81 |
(fn t => "Conclusion of rule does not name a recursive set: " ^ |
82 |
Sign.string_of_term sign t); |
|
923 | 83 |
end; |
84 |
||
85 |
val _ = assert_all is_Free rec_params |
|
1465 | 86 |
(fn t => "Param in recursion term not a free variable: " ^ |
87 |
Sign.string_of_term sign t); |
|
923 | 88 |
|
89 |
(*** Construct the lfp definition ***) |
|
90 |
val mk_variant = variant (foldr add_term_names (intr_tms,[])); |
|
91 |
||
92 |
val z = mk_variant"z" and X = mk_variant"X" and w = mk_variant"w"; |
|
93 |
||
1189 | 94 |
(*Mutual recursion ?? *) |
2995
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
95 |
val dom_set = body_type recT |
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
96 |
val dom_sumT = dest_setT dom_set |
923 | 97 |
|
98 |
val freez = Free(z, dom_sumT) |
|
99 |
and freeX = Free(X, dom_set); |
|
100 |
||
101 |
fun dest_tprop (Const("Trueprop",_) $ P) = P |
|
102 |
| dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ |
|
1465 | 103 |
Sign.string_of_term sign Q); |
923 | 104 |
|
105 |
(*Makes a disjunct from an introduction rule*) |
|
106 |
fun lfp_part intr = (*quantify over rule's free vars except parameters*) |
|
1397 | 107 |
let val prems = map dest_tprop (Logic.strip_imp_prems intr) |
1465 | 108 |
val _ = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds |
109 |
val exfrees = term_frees intr \\ rec_params |
|
110 |
val zeq = eq_const dom_sumT $ freez $ (#1 (rule_concl intr)) |
|
923 | 111 |
in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end; |
112 |
||
113 |
(*The Part(A,h) terms -- compose injections to make h*) |
|
1465 | 114 |
fun mk_Part (Bound 0, _) = freeX (*no mutual rec, no Part needed*) |
923 | 115 |
| mk_Part (h, domT) = |
1465 | 116 |
let val goodh = mend_sum_types (h, dom_sumT) |
923 | 117 |
and Part_const = |
1465 | 118 |
Const("Part", [dom_set, domT-->dom_sumT]---> dom_set) |
923 | 119 |
in Part_const $ freeX $ Abs(w,domT,goodh) end; |
120 |
||
2995
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
121 |
(*Access to balanced disjoint sums via injections?? |
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
122 |
Mutual recursion is NOT supported*) |
2270 | 123 |
val parts = ListPair.map mk_Part |
2995
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
124 |
(accesses_bal (ap Inl, ap Inr, Bound 0) 1, |
84df3b150b67
Disabled the attempts for mutual induction to work so that single induction
paulson
parents:
2859
diff
changeset
|
125 |
[dom_set]); |
923 | 126 |
|
127 |
(*replace each set by the corresponding Part(A,h)*) |
|
128 |
val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms; |
|
129 |
||
130 |
val lfp_rhs = Fp.oper(X, dom_sumT, |
|
1465 | 131 |
mk_Collect(z, dom_sumT, |
132 |
fold_bal (app disj) part_intrs)) |
|
923 | 133 |
|
134 |
||
135 |
(*** Make the new theory ***) |
|
136 |
||
137 |
(*A key definition: |
|
138 |
If no mutual recursion then it equals the one recursive set. |
|
139 |
If mutual recursion then it differs from all the recursive sets. *) |
|
3945 | 140 |
val big_rec_base_name = space_implode "_" rec_base_names; |
4235 | 141 |
val big_rec_name = Sign.intern_const sign big_rec_base_name; |
923 | 142 |
|
143 |
(*Big_rec... is the union of the mutually recursive sets*) |
|
144 |
val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); |
|
145 |
||
146 |
(*The individual sets must already be declared*) |
|
4871 | 147 |
val axpairs = map Logic.mk_defpair |
1465 | 148 |
((big_rec_tm, lfp_rhs) :: |
149 |
(case parts of |
|
150 |
[_] => [] (*no mutual recursion*) |
|
151 |
| _ => rec_tms ~~ (*define the sets as Parts*) |
|
152 |
map (subst_atomic [(freeX, big_rec_tm)]) parts)); |
|
923 | 153 |
|
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4235
diff
changeset
|
154 |
(*tracing: print the fixedpoint definition*) |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4235
diff
changeset
|
155 |
val _ = if !Ind_Syntax.trace then |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4235
diff
changeset
|
156 |
seq (writeln o Sign.string_of_term sign o #2) axpairs |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4235
diff
changeset
|
157 |
else () |
923 | 158 |
|
1397 | 159 |
(*Detect occurrences of operator, even with other types!*) |
160 |
val _ = (case rec_names inter (add_term_names (lfp_rhs,[])) of |
|
1465 | 161 |
[] => () |
162 |
| x::_ => error ("Illegal occurrence of recursion op: " ^ x ^ |
|
1397 | 163 |
"\n\t*Consider adding type constraints*")) |
164 |
||
4871 | 165 |
in thy |> PureThy.add_defs_i (map Attribute.none axpairs) end |
923 | 166 |
|
167 |
||
168 |
end; |