author | nipkow |
Wed, 09 Nov 1994 19:51:09 +0100 | |
changeset 168 | 44ff2275d44f |
parent 134 | 4b7da5a895e7 |
child 182 | d5c6d1fb236b |
permissions | -rw-r--r-- |
128 | 1 |
(* Title: HOL/ind_syntax.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1994 University of Cambridge |
|
5 |
||
6 |
Abstract Syntax functions for Inductive Definitions |
|
7 |
See also ../Pure/section-utils.ML |
|
8 |
*) |
|
9 |
||
10 |
(*The structure protects these items from redeclaration (somewhat!). The |
|
11 |
datatype definitions in theory files refer to these items by name! |
|
12 |
*) |
|
13 |
structure Ind_Syntax = |
|
14 |
struct |
|
15 |
||
16 |
(** Abstract syntax definitions for HOL **) |
|
17 |
||
18 |
val termC: class = "term"; |
|
19 |
val termS: sort = [termC]; |
|
20 |
||
21 |
val termTVar = TVar(("'a",0),termS); |
|
22 |
||
23 |
val boolT = Type("bool",[]); |
|
24 |
val unitT = Type("unit",[]); |
|
25 |
||
26 |
val conj = Const("op &", [boolT,boolT]--->boolT) |
|
27 |
and disj = Const("op |", [boolT,boolT]--->boolT) |
|
28 |
and imp = Const("op -->", [boolT,boolT]--->boolT); |
|
29 |
||
30 |
fun eq_const T = Const("op =", [T,T]--->boolT); |
|
31 |
||
32 |
fun mk_set T = Type("set",[T]); |
|
33 |
||
34 |
fun dest_set (Type("set",[T])) = T |
|
35 |
| dest_set _ = error "dest_set: set type expected" |
|
36 |
||
37 |
fun mk_mem (x,A) = |
|
38 |
let val setT = fastype_of A |
|
39 |
in Const("op :", [dest_set setT, setT]--->boolT) $ x $ A |
|
40 |
end; |
|
41 |
||
42 |
fun Int_const T = |
|
43 |
let val sT = mk_set T |
|
44 |
in Const("op Int", [sT,sT]--->sT) end; |
|
45 |
||
46 |
fun exists_const T = Const("Ex", [T-->boolT]--->boolT); |
|
47 |
fun mk_exists (Free(x,T),P) = exists_const T $ (absfree (x,T,P)); |
|
48 |
||
49 |
fun all_const T = Const("All", [T-->boolT]--->boolT); |
|
50 |
fun mk_all (Free(x,T),P) = all_const T $ (absfree (x,T,P)); |
|
51 |
||
52 |
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *) |
|
53 |
fun mk_all_imp (A,P) = |
|
54 |
let val T = dest_set (fastype_of A) |
|
55 |
in all_const T $ Abs("v", T, imp $ (mk_mem (Bound 0, A)) $ (P $ Bound 0)) |
|
56 |
end; |
|
57 |
||
58 |
(** Cartesian product type **) |
|
59 |
||
60 |
fun mk_prod (T1,T2) = Type("*", [T1,T2]); |
|
61 |
||
134
4b7da5a895e7
HOL/ind_syntax/factors: now returns only factors in the product type that
lcp
parents:
128
diff
changeset
|
62 |
(*Maps the type T1*...*Tn to [T1,...,Tn], if nested to the right*) |
4b7da5a895e7
HOL/ind_syntax/factors: now returns only factors in the product type that
lcp
parents:
128
diff
changeset
|
63 |
fun factors (Type("*", [T1,T2])) = T1 :: factors T2 |
128 | 64 |
| factors T = [T]; |
65 |
||
66 |
(*Make a correctly typed ordered pair*) |
|
67 |
fun mk_Pair (t1,t2) = |
|
68 |
let val T1 = fastype_of t1 |
|
69 |
and T2 = fastype_of t2 |
|
70 |
in Const("Pair", [T1, T2] ---> mk_prod(T1,T2)) $ t1 $ t2 end; |
|
71 |
||
72 |
fun split_const(Ta,Tb,Tc) = |
|
73 |
Const("split", [[Ta,Tb]--->Tc, mk_prod(Ta,Tb)] ---> Tc); |
|
74 |
||
75 |
(*Given u expecting arguments of types [T1,...,Tn], create term of |
|
76 |
type T1*...*Tn => Tc using split. Here * associates to the LEFT*) |
|
77 |
fun ap_split_l Tc u [ ] = Abs("null", unitT, u) |
|
78 |
| ap_split_l Tc u [_] = u |
|
79 |
| ap_split_l Tc u (Ta::Tb::Ts) = ap_split_l Tc (split_const(Ta,Tb,Tc) $ u) |
|
80 |
(mk_prod(Ta,Tb) :: Ts); |
|
81 |
||
82 |
(*Given u expecting arguments of types [T1,...,Tn], create term of |
|
83 |
type T1*...*Tn => i using split. Here * associates to the RIGHT*) |
|
84 |
fun ap_split Tc u [ ] = Abs("null", unitT, u) |
|
85 |
| ap_split Tc u [_] = u |
|
86 |
| ap_split Tc u [Ta,Tb] = split_const(Ta,Tb,Tc) $ u |
|
87 |
| ap_split Tc u (Ta::Ts) = |
|
88 |
split_const(Ta, foldr1 mk_prod Ts, Tc) $ |
|
89 |
(Abs("v", Ta, ap_split Tc (u $ Bound(length Ts - 2)) Ts)); |
|
90 |
||
91 |
(** Disjoint sum type **) |
|
92 |
||
93 |
fun mk_sum (T1,T2) = Type("+", [T1,T2]); |
|
94 |
val Inl = Const("Inl", dummyT) |
|
95 |
and Inr = Const("Inr", dummyT); (*correct types added later!*) |
|
96 |
(*val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT)*) |
|
97 |
||
98 |
fun summands (Type("+", [T1,T2])) = summands T1 @ summands T2 |
|
99 |
| summands T = [T]; |
|
100 |
||
101 |
(*Given the destination type, fills in correct types of an Inl/Inr nest*) |
|
102 |
fun mend_sum_types (h,T) = |
|
103 |
(case (h,T) of |
|
104 |
(Const("Inl",_) $ h1, Type("+", [T1,T2])) => |
|
105 |
Const("Inl", T1 --> T) $ (mend_sum_types (h1, T1)) |
|
106 |
| (Const("Inr",_) $ h2, Type("+", [T1,T2])) => |
|
107 |
Const("Inr", T2 --> T) $ (mend_sum_types (h2, T2)) |
|
108 |
| _ => h); |
|
109 |
||
110 |
fun Collect_const T = Const("Collect", [T-->boolT] ---> mk_set T); |
|
111 |
fun mk_Collect (a,T,t) = Collect_const T $ absfree(a, T, t); |
|
112 |
||
113 |
val Trueprop = Const("Trueprop",boolT-->propT); |
|
114 |
fun mk_tprop P = Trueprop $ P; |
|
115 |
||
116 |
||
117 |
(*simple error-checking in the premises of an inductive definition*) |
|
118 |
fun chk_prem rec_hd (Const("op &",_) $ _ $ _) = |
|
119 |
error"Premises may not be conjuctive" |
|
120 |
| chk_prem rec_hd (Const("op :",_) $ t $ X) = |
|
121 |
deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol" |
|
122 |
| chk_prem rec_hd t = |
|
123 |
deny (Logic.occs(rec_hd,t)) "Recursion term in side formula"; |
|
124 |
||
125 |
(*Return the conclusion of a rule, of the form t:X*) |
|
126 |
fun rule_concl rl = |
|
127 |
let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = |
|
128 |
Logic.strip_imp_concl rl |
|
129 |
in (t,X) end; |
|
130 |
||
131 |
(*As above, but return error message if bad*) |
|
132 |
fun rule_concl_msg sign rl = rule_concl rl |
|
133 |
handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ |
|
134 |
Sign.string_of_term sign rl); |
|
135 |
||
136 |
(*For simplifying the elimination rule*) |
|
137 |
val sumprod_free_SEs = |
|
138 |
Pair_inject :: |
|
139 |
map make_elim [Inl_neq_Inr, Inr_neq_Inl, Inl_inject, Inr_inject]; |
|
140 |
||
141 |
(*For deriving cases rules. |
|
142 |
read_instantiate replaces a propositional variable by a formula variable*) |
|
143 |
val equals_CollectD = |
|
144 |
read_instantiate [("W","?Q")] |
|
145 |
(make_elim (equalityD1 RS subsetD RS CollectD)); |
|
146 |
||
147 |
(*Delete needless equality assumptions*) |
|
148 |
val refl_thin = prove_goal HOL.thy "!!P. [| a=a; P |] ==> P" |
|
149 |
(fn _ => [assume_tac 1]); |
|
150 |
||
151 |
end; |
|
152 |