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