| author | wenzelm | 
| Fri, 16 May 2025 12:10:49 +0200 | |
| changeset 82626 | e840461d5370 | 
| parent 74375 | ba880f3a4e52 | 
| permissions | -rw-r--r-- | 
| 12243 | 1  | 
(* Title: ZF/ind_syntax.ML  | 
| 1461 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 3  | 
Copyright 1993 University of Cambridge  | 
4  | 
||
| 12243 | 5  | 
Abstract Syntax functions for Inductive Definitions.  | 
| 0 | 6  | 
*)  | 
7  | 
||
| 516 | 8  | 
structure Ind_Syntax =  | 
9  | 
struct  | 
|
| 0 | 10  | 
|
| 
4804
 
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
 
paulson 
parents: 
4352 
diff
changeset
 | 
11  | 
(*Print tracing messages during processing of "inductive" theory sections*)  | 
| 32740 | 12  | 
val trace = Unsynchronized.ref false;  | 
| 
4804
 
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
 
paulson 
parents: 
4352 
diff
changeset
 | 
13  | 
|
| 26189 | 14  | 
fun traceIt msg thy t =  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26189 
diff
changeset
 | 
15  | 
if !trace then (tracing (msg ^ Syntax.string_of_term_global thy t); t)  | 
| 17988 | 16  | 
else t;  | 
17  | 
||
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
18  | 
|
| 
4352
 
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
 
paulson 
parents: 
3925 
diff
changeset
 | 
19  | 
(** Abstract syntax definitions for ZF **)  | 
| 0 | 20  | 
|
21  | 
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)  | 
|
| 26189 | 22  | 
fun mk_all_imp (A,P) =  | 
| 74320 | 23  | 
let val T = \<^Type>\<open>i\<close> in  | 
24  | 
\<^Const>\<open>All T for  | 
|
| 74375 | 25  | 
      \<open>Abs ("v", T, \<^Const>\<open>imp for \<^Const>\<open>mem for \<open>Bound 0\<close> A\<close> \<open>Term.betapply (P, Bound 0)\<close>\<close>)\<close>\<close>
 | 
| 74320 | 26  | 
end;  | 
| 0 | 27  | 
|
| 
74319
 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 
wenzelm 
parents: 
74296 
diff
changeset
 | 
28  | 
fun mk_Collect (a, D, t) = \<^Const>\<open>Collect for D\<close> $ absfree (a, \<^Type>\<open>i\<close>) t;  | 
| 0 | 29  | 
|
| 516 | 30  | 
(*simple error-checking in the premises of an inductive definition*)  | 
| 74294 | 31  | 
fun chk_prem rec_hd \<^Const_>\<open>conj for _ _\<close> =  | 
| 1461 | 32  | 
error"Premises may not be conjuctive"  | 
| 74294 | 33  | 
| chk_prem rec_hd \<^Const_>\<open>mem for t X\<close> =  | 
| 
22567
 
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
 
wenzelm 
parents: 
21539 
diff
changeset
 | 
34  | 
(Logic.occs(rec_hd,t) andalso error "Recursion term on left of member symbol"; ())  | 
| 26189 | 35  | 
| 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
 | 
36  | 
(Logic.occs(rec_hd,t) andalso error "Recursion term in side formula"; ());  | 
| 516 | 37  | 
|
| 
14
 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
6 
diff
changeset
 | 
38  | 
(*Return the conclusion of a rule, of the form t:X*)  | 
| 26189 | 39  | 
fun rule_concl rl =  | 
| 74375 | 40  | 
let val \<^Const_>\<open>Trueprop for \<^Const_>\<open>mem for t X\<close>\<close> = Logic.strip_imp_concl rl  | 
| 435 | 41  | 
in (t,X) end;  | 
42  | 
||
43  | 
(*As above, but return error message if bad*)  | 
|
44  | 
fun rule_concl_msg sign rl = rule_concl rl  | 
|
| 26189 | 45  | 
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26189 
diff
changeset
 | 
46  | 
Syntax.string_of_term_global sign rl);  | 
| 0 | 47  | 
|
48  | 
(*For deriving cases rules. CollectD2 discards the domain, which is redundant;  | 
|
49  | 
read_instantiate replaces a propositional variable by a formula variable*)  | 
|
| 26189 | 50  | 
val equals_CollectD =  | 
| 69593 | 51  | 
    Rule_Insts.read_instantiate \<^context> [((("W", 0), Position.none), "Q")] ["Q"]
 | 
| 24893 | 52  | 
        (make_elim (@{thm equalityD1} RS @{thm subsetD} RS @{thm CollectD2}));
 | 
| 0 | 53  | 
|
54  | 
||
| 516 | 55  | 
(** For datatype definitions **)  | 
56  | 
||
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
57  | 
(*Constructor name, type, mixfix info;  | 
| 
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
58  | 
internal name from mixfix, datatype sets, full premises*)  | 
| 26189 | 59  | 
type constructor_spec =  | 
60  | 
(string * typ * mixfix) * string * term list * term list;  | 
|
| 
6053
 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 
paulson 
parents: 
4972 
diff
changeset
 | 
61  | 
|
| 74294 | 62  | 
fun dest_mem \<^Const_>\<open>mem for x A\<close> = (x, A)  | 
| 516 | 63  | 
| dest_mem _ = error "Constructor specifications must have the form x:A";  | 
64  | 
||
65  | 
(*read a constructor specification*)  | 
|
| 35129 | 66  | 
fun read_construct ctxt (id: string, sprems, syn: mixfix) =  | 
| 
74319
 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 
wenzelm 
parents: 
74296 
diff
changeset
 | 
67  | 
let val prems = map (Syntax.parse_term ctxt #> Type.constraint \<^Type>\<open>o\<close>) sprems  | 
| 
27261
 
5b3101338f42
eliminated old Sign.read_term/Thm.read_cterm etc.;
 
wenzelm 
parents: 
27239 
diff
changeset
 | 
68  | 
|> Syntax.check_terms ctxt  | 
| 1461 | 69  | 
val args = map (#1 o dest_mem) prems  | 
| 
74319
 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 
wenzelm 
parents: 
74296 
diff
changeset
 | 
70  | 
val T = (map (#2 o dest_Free) args) ---> \<^Type>\<open>i\<close>  | 
| 26189 | 71  | 
handle TERM _ => error  | 
| 1461 | 72  | 
"Bad variable in constructor specification"  | 
| 35129 | 73  | 
in ((id,T,syn), id, args, prems) end;  | 
| 516 | 74  | 
|
75  | 
val read_constructs = map o map o read_construct;  | 
|
| 0 | 76  | 
|
| 516 | 77  | 
(*convert constructor specifications into introduction rules*)  | 
| 
3925
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
78  | 
fun mk_intr_tms sg (rec_tm, constructs) =  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
79  | 
let  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
80  | 
fun mk_intr ((id,T,syn), name, args, prems) =  | 
| 
 
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
 
wenzelm 
parents: 
2266 
diff
changeset
 | 
81  | 
Logic.list_implies  | 
| 74342 | 82  | 
(map \<^make_judgment> prems,  | 
83  | 
\<^make_judgment> (\<^Const>\<open>mem\<close> $ list_comb (Const (Sign.full_bname sg name, T), args) $ rec_tm))  | 
|
| 516 | 84  | 
in map mk_intr constructs end;  | 
85  | 
||
| 32952 | 86  | 
fun mk_all_intr_tms sg arg = flat (ListPair.map (mk_intr_tms sg) arg);  | 
| 0 | 87  | 
|
| 74296 | 88  | 
fun mk_Un (t1, t2) = \<^Const>\<open>Un for t1 t2\<close>;  | 
| 0 | 89  | 
|
| 516 | 90  | 
(*Make a datatype's domain: form the union of its set parameters*)  | 
| 6112 | 91  | 
fun union_params (rec_tm, cs) =  | 
| 516 | 92  | 
let val (_,args) = strip_comb rec_tm  | 
| 
74319
 
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
 
wenzelm 
parents: 
74296 
diff
changeset
 | 
93  | 
fun is_ind arg = (type_of arg = \<^Type>\<open>i\<close>)  | 
| 33317 | 94  | 
in case filter is_ind (args @ cs) of  | 
| 74296 | 95  | 
[] => \<^Const>\<open>zero\<close>  | 
| 32765 | 96  | 
| u_args => Balanced_Tree.make mk_Un u_args  | 
| 516 | 97  | 
end;  | 
98  | 
||
| 0 | 99  | 
|
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
100  | 
(*Includes rules for succ and Pair since they are common constructions*)  | 
| 26189 | 101  | 
val elim_rls =  | 
102  | 
  [@{thm asm_rl}, @{thm FalseE}, @{thm succ_neq_0}, @{thm sym} RS @{thm succ_neq_0},
 | 
|
103  | 
   @{thm Pair_neq_0}, @{thm sym} RS @{thm Pair_neq_0}, @{thm Pair_inject},
 | 
|
104  | 
   make_elim @{thm succ_inject}, @{thm refl_thin}, @{thm conjE}, @{thm exE}, @{thm disjE}];
 | 
|
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
105  | 
|
| 7694 | 106  | 
|
107  | 
(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)  | 
|
108  | 
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))  | 
|
109  | 
  | tryres (th, []) = raise THM("tryres", 0, [th]);
 | 
|
110  | 
||
| 26189 | 111  | 
fun gen_make_elim elim_rls rl =  | 
| 
35021
 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 
wenzelm 
parents: 
33317 
diff
changeset
 | 
112  | 
Drule.export_without_context (tryres (rl, elim_rls @ [revcut_rl]));  | 
| 7694 | 113  | 
|
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
114  | 
(*Turns iff rules into safe elimination rules*)  | 
| 26189 | 115  | 
fun mk_free_SEs iffs = map (gen_make_elim [@{thm conjE}, @{thm FalseE}]) (iffs RL [@{thm iffD1}]);
 | 
| 
1418
 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 
paulson 
parents: 
742 
diff
changeset
 | 
116  | 
|
| 516 | 117  | 
end;  | 
118  |