| author | haftmann | 
| Tue, 21 Aug 2007 09:07:04 +0200 | |
| changeset 24379 | 823ffe1fdf67 | 
| parent 23419 | 8c30dd4b3b22 | 
| child 24826 | 78e6a3cea367 | 
| permissions | -rw-r--r-- | 
| 12243 | 1  | 
(* Title: ZF/ind_syntax.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
| 12243 | 6  | 
Abstract Syntax functions for Inductive Definitions.  | 
| 0 | 7  | 
*)  | 
8  | 
||
| 516 | 9  | 
(*The structure protects these items from redeclaration (somewhat!). The  | 
10  | 
datatype definitions in theory files refer to these items by name!  | 
|
11  | 
*)  | 
|
12  | 
structure Ind_Syntax =  | 
|
13  | 
struct  | 
|
| 0 | 14  | 
|
| 
4804
 
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
 
paulson 
parents: 
4352 
diff
changeset
 | 
15  | 
(*Print tracing messages during processing of "inductive" theory sections*)  | 
| 
 
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
 
paulson 
parents: 
4352 
diff
changeset
 | 
16  | 
val trace = ref false;  | 
| 
 
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
 
paulson 
parents: 
4352 
diff
changeset
 | 
17  | 
|
| 17988 | 18  | 
fun traceIt msg thy t =  | 
19  | 
if !trace then (tracing (msg ^ Sign.string_of_term thy t); t)  | 
|
20  | 
else t;  | 
|
21  | 
||
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
22  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
23  | 
(** Abstract syntax definitions for ZF **)  | 
| 0 | 24  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
25  | 
val iT = Type("i",[]);
 | 
| 0 | 26  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
27  | 
val mem_const = Const("op :", [iT,iT]--->FOLogic.oT);
 | 
| 0 | 28  | 
|
29  | 
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)  | 
|
30  | 
fun mk_all_imp (A,P) =  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
31  | 
FOLogic.all_const iT $  | 
| 
4972
 
7fe1d30c1374
mk_all_imp: no longer creates goals that have beta-redexes
 
paulson 
parents: 
4804 
diff
changeset
 | 
32  | 
      Abs("v", iT, FOLogic.imp $ (mem_const $ Bound 0 $ A) $ 
 | 
| 18176 | 33  | 
Term.betapply(P, Bound 0));  | 
| 0 | 34  | 
|
35  | 
val Part_const = Const("Part", [iT,iT-->iT]--->iT);
 | 
|
36  | 
||
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
37  | 
val apply_const = Const("op `", [iT,iT]--->iT);
 | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
38  | 
|
| 6093 | 39  | 
val Vrecursor_const = Const("Univ.Vrecursor", [[iT,iT]--->iT, iT]--->iT);
 | 
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
40  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
41  | 
val Collect_const = Const("Collect", [iT, iT-->FOLogic.oT] ---> iT);
 | 
| 0 | 42  | 
fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t);  | 
43  | 
||
| 516 | 44  | 
(*simple error-checking in the premises of an inductive definition*)  | 
45  | 
fun chk_prem rec_hd (Const("op &",_) $ _ $ _) =
 | 
|
| 1461 | 46  | 
error"Premises may not be conjuctive"  | 
| 516 | 47  | 
  | chk_prem rec_hd (Const("op :",_) $ t $ X) = 
 | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
21539 
diff
changeset
 | 
48  | 
(Logic.occs(rec_hd,t) andalso error "Recursion term on left of member symbol"; ())  | 
| 516 | 49  | 
| chk_prem rec_hd t =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
21539 
diff
changeset
 | 
50  | 
(Logic.occs(rec_hd,t) andalso error "Recursion term in side formula"; ());  | 
| 516 | 51  | 
|
| 
14
 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
6 
diff
changeset
 | 
52  | 
(*Return the conclusion of a rule, of the form t:X*)  | 
| 0 | 53  | 
fun rule_concl rl =  | 
| 435 | 54  | 
    let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = 
 | 
| 1461 | 55  | 
Logic.strip_imp_concl rl  | 
| 435 | 56  | 
in (t,X) end;  | 
57  | 
||
58  | 
(*As above, but return error message if bad*)  | 
|
59  | 
fun rule_concl_msg sign rl = rule_concl rl  | 
|
60  | 
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ 
 | 
|
| 1461 | 61  | 
Sign.string_of_term sign rl);  | 
| 0 | 62  | 
|
63  | 
(*For deriving cases rules. CollectD2 discards the domain, which is redundant;  | 
|
64  | 
read_instantiate replaces a propositional variable by a formula variable*)  | 
|
65  | 
val equals_CollectD =  | 
|
66  | 
    read_instantiate [("W","?Q")]
 | 
|
67  | 
(make_elim (equalityD1 RS subsetD RS CollectD2));  | 
|
68  | 
||
69  | 
||
| 516 | 70  | 
(** For datatype definitions **)  | 
71  | 
||
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
72  | 
(*Constructor name, type, mixfix info;  | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
73  | 
internal name from mixfix, datatype sets, full premises*)  | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
74  | 
type constructor_spec =  | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
75  | 
((string * typ * mixfix) * string * term list * term list);  | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
76  | 
|
| 516 | 77  | 
fun dest_mem (Const("op :",_) $ x $ A) = (x,A)
 | 
78  | 
| dest_mem _ = error "Constructor specifications must have the form x:A";  | 
|
79  | 
||
80  | 
(*read a constructor specification*)  | 
|
81  | 
fun read_construct sign (id, sprems, syn) =  | 
|
| 8819 | 82  | 
let val prems = map (Sign.simple_read_term sign FOLogic.oT) sprems  | 
| 1461 | 83  | 
val args = map (#1 o dest_mem) prems  | 
84  | 
val T = (map (#2 o dest_Free) args) ---> iT  | 
|
85  | 
handle TERM _ => error  | 
|
86  | 
"Bad variable in constructor specification"  | 
|
| 568 | 87  | 
val name = Syntax.const_name id syn (*handle infix constructors*)  | 
| 516 | 88  | 
in ((id,T,syn), name, args, prems) end;  | 
89  | 
||
90  | 
val read_constructs = map o map o read_construct;  | 
|
| 0 | 91  | 
|
| 516 | 92  | 
(*convert constructor specifications into introduction rules*)  | 
| 
3925
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
93  | 
fun mk_intr_tms sg (rec_tm, constructs) =  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
94  | 
let  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
95  | 
fun mk_intr ((id,T,syn), name, args, prems) =  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
96  | 
Logic.list_implies  | 
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
97  | 
(map FOLogic.mk_Trueprop prems,  | 
| 
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
98  | 
FOLogic.mk_Trueprop  | 
| 
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
99  | 
(mem_const $ list_comb (Const (Sign.full_name sg name, T), args)  | 
| 
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
100  | 
$ rec_tm))  | 
| 516 | 101  | 
in map mk_intr constructs end;  | 
102  | 
||
| 
3925
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
103  | 
fun mk_all_intr_tms sg arg = List.concat (ListPair.map (mk_intr_tms sg) arg);  | 
| 0 | 104  | 
|
| 7694 | 105  | 
fun mk_Un (t1, t2) = Const("op Un", [iT,iT]--->iT) $ t1 $ t2;
 | 
106  | 
||
107  | 
val empty       = Const("0", iT)
 | 
|
| 6093 | 108  | 
and univ        = Const("Univ.univ", iT-->iT)
 | 
109  | 
and quniv       = Const("QUniv.quniv", iT-->iT);
 | 
|
| 0 | 110  | 
|
| 516 | 111  | 
(*Make a datatype's domain: form the union of its set parameters*)  | 
| 6112 | 112  | 
fun union_params (rec_tm, cs) =  | 
| 516 | 113  | 
let val (_,args) = strip_comb rec_tm  | 
| 6112 | 114  | 
fun is_ind arg = (type_of arg = iT)  | 
| 15570 | 115  | 
in case List.filter is_ind (args @ cs) of  | 
| 6112 | 116  | 
[] => empty  | 
| 23419 | 117  | 
| u_args => BalancedTree.make mk_Un u_args  | 
| 516 | 118  | 
end;  | 
119  | 
||
| 6112 | 120  | 
(*univ or quniv constitutes the sum domain for mutual recursion;  | 
121  | 
it is applied to the datatype parameters and to Consts occurring in the  | 
|
122  | 
definition other than Nat.nat and the datatype sets themselves.  | 
|
123  | 
FIXME: could insert all constant set expressions, e.g. nat->nat.*)  | 
|
124  | 
fun data_domain co (rec_tms, con_ty_lists) =  | 
|
| 
13150
 
0c50d13d449a
better error messages for datatypes not declared Const
 
paulson 
parents: 
12243 
diff
changeset
 | 
125  | 
let val rec_hds = map head_of rec_tms  | 
| 
 
0c50d13d449a
better error messages for datatypes not declared Const
 
paulson 
parents: 
12243 
diff
changeset
 | 
126  | 
val dummy = assert_all is_Const rec_hds  | 
| 
 
0c50d13d449a
better error messages for datatypes not declared Const
 
paulson 
parents: 
12243 
diff
changeset
 | 
127  | 
(fn t => "Datatype set not previously declared as constant: " ^  | 
| 23156 | 128  | 
                   Sign.string_of_term @{theory IFOL} t);
 | 
| 
13150
 
0c50d13d449a
better error messages for datatypes not declared Const
 
paulson 
parents: 
12243 
diff
changeset
 | 
129  | 
val rec_names = (*nat doesn't have to be added*)  | 
| 
 
0c50d13d449a
better error messages for datatypes not declared Const
 
paulson 
parents: 
12243 
diff
changeset
 | 
130  | 
"Nat.nat" :: map (#1 o dest_Const) rec_hds  | 
| 6112 | 131  | 
val u = if co then quniv else univ  | 
| 16867 | 132  | 
val cs = (fold o fold) (fn (_, _, _, prems) => prems |> (fold o fold_aterms)  | 
133  | 
(fn t as Const (a, _) => if a mem_string rec_names then I else insert (op =) t  | 
|
134  | 
| _ => I)) con_ty_lists [];  | 
|
| 6112 | 135  | 
in u $ union_params (hd rec_tms, cs) end;  | 
136  | 
||
| 0 | 137  | 
|
138  | 
(*Could go to FOL, but it's hardly general*)  | 
|
| 21539 | 139  | 
val def_swap_iff = prove_goal (the_context ()) "a==b ==> a=c <-> c=b"  | 
140  | 
(fn [def] => [(rewtac def), (rtac iffI 1), (REPEAT (etac sym 1))]);  | 
|
| 0 | 141  | 
|
| 21539 | 142  | 
val def_trans = prove_goal (the_context ()) "[| f==g; g(a)=b |] ==> f(a)=b"  | 
| 0 | 143  | 
(fn [rew,prem] => [ rewtac rew, rtac prem 1 ]);  | 
144  | 
||
| 55 | 145  | 
(*Delete needless equality assumptions*)  | 
| 21539 | 146  | 
val refl_thin = prove_goal (the_context ()) "!!P. [| a=a; P |] ==> P"  | 
| 55 | 147  | 
(fn _ => [assume_tac 1]);  | 
| 0 | 148  | 
|
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
149  | 
(*Includes rules for succ and Pair since they are common constructions*)  | 
| 21539 | 150  | 
val elim_rls = [asm_rl, FalseE, thm "succ_neq_0", sym RS thm "succ_neq_0",  | 
151  | 
thm "Pair_neq_0", sym RS thm "Pair_neq_0", thm "Pair_inject",  | 
|
152  | 
make_elim (thm "succ_inject"),  | 
|
| 1461 | 153  | 
refl_thin, conjE, exE, disjE];  | 
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
154  | 
|
| 7694 | 155  | 
|
156  | 
(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)  | 
|
157  | 
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))  | 
|
158  | 
  | tryres (th, []) = raise THM("tryres", 0, [th]);
 | 
|
159  | 
||
160  | 
fun gen_make_elim elim_rls rl =  | 
|
161  | 
standard (tryres (rl, elim_rls @ [revcut_rl]));  | 
|
162  | 
||
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
163  | 
(*Turns iff rules into safe elimination rules*)  | 
| 
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
164  | 
fun mk_free_SEs iffs = map (gen_make_elim [conjE,FalseE]) (iffs RL [iffD1]);  | 
| 
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
165  | 
|
| 516 | 166  | 
end;  | 
167  | 
||
| 6112 | 168  | 
|
169  | 
(*For convenient access by the user*)  | 
|
170  | 
val trace_induct = Ind_Syntax.trace;  |