| 0 |      1 | (*  Title:      CTT/ctt.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1993  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Constructive Type Theory
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | CTT = Pure +
 | 
|  |     10 | 
 | 
| 283 |     11 | types
 | 
|  |     12 |   i
 | 
|  |     13 |   t
 | 
|  |     14 |   o
 | 
| 0 |     15 | 
 | 
| 283 |     16 | arities
 | 
|  |     17 |    i,t,o :: logic
 | 
| 0 |     18 | 
 | 
|  |     19 | consts
 | 
|  |     20 |   (*Types*)
 | 
|  |     21 |   F,T       :: "t"          (*F is empty, T contains one element*)
 | 
|  |     22 |   contr     :: "i=>i"
 | 
|  |     23 |   tt        :: "i"
 | 
|  |     24 |   (*Natural numbers*)
 | 
|  |     25 |   N         :: "t"
 | 
|  |     26 |   succ      :: "i=>i"
 | 
|  |     27 |   rec       :: "[i, i, [i,i]=>i] => i"
 | 
|  |     28 |   (*Unions*)
 | 
|  |     29 |   inl,inr   :: "i=>i"
 | 
|  |     30 |   when      :: "[i, i=>i, i=>i]=>i"
 | 
|  |     31 |   (*General Sum and Binary Product*)
 | 
|  |     32 |   Sum       :: "[t, i=>t]=>t"
 | 
|  |     33 |   fst,snd   :: "i=>i"
 | 
|  |     34 |   split     :: "[i, [i,i]=>i] =>i"
 | 
|  |     35 |   (*General Product and Function Space*)
 | 
|  |     36 |   Prod      :: "[t, i=>t]=>t"
 | 
|  |     37 |   (*Equality type*)
 | 
|  |     38 |   Eq        :: "[t,i,i]=>t"
 | 
|  |     39 |   eq        :: "i"
 | 
|  |     40 |   (*Judgements*)
 | 
|  |     41 |   Type      :: "t => prop"          ("(_ type)" [10] 5)
 | 
|  |     42 |   Eqtype    :: "[t,t]=>prop"        ("(3_ =/ _)" [10,10] 5)
 | 
|  |     43 |   Elem      :: "[i, t]=>prop"       ("(_ /: _)" [10,10] 5)
 | 
|  |     44 |   Eqelem    :: "[i,i,t]=>prop"      ("(3_ =/ _ :/ _)" [10,10,10] 5)
 | 
|  |     45 |   Reduce    :: "[i,i]=>prop"        ("Reduce[_,_]")
 | 
|  |     46 |   (*Types*)
 | 
| 23 |     47 |   "@PROD"   :: "[idt,t,t]=>t"       ("(3PROD _:_./ _)" 10)
 | 
|  |     48 |   "@SUM"    :: "[idt,t,t]=>t"       ("(3SUM _:_./ _)" 10)
 | 
| 0 |     49 |   "+"       :: "[t,t]=>t"           (infixr 40)
 | 
|  |     50 |   (*Invisible infixes!*)
 | 
|  |     51 |   "@-->"    :: "[t,t]=>t"           ("(_ -->/ _)" [31,30] 30)
 | 
|  |     52 |   "@*"      :: "[t,t]=>t"           ("(_ */ _)" [51,50] 50)
 | 
|  |     53 |   (*Functions*)
 | 
|  |     54 |   lambda    :: "(i => i) => i"      (binder "lam " 10)
 | 
|  |     55 |   "`"       :: "[i,i]=>i"           (infixl 60)
 | 
|  |     56 |   (*Natural numbers*)
 | 
|  |     57 |   "0"       :: "i"                  ("0")
 | 
|  |     58 |   (*Pairing*)
 | 
|  |     59 |   pair      :: "[i,i]=>i"           ("(1<_,/_>)")
 | 
|  |     60 | 
 | 
|  |     61 | translations
 | 
|  |     62 |   "PROD x:A. B" => "Prod(A, %x. B)"
 | 
| 23 |     63 |   "A --> B"     => "Prod(A, _K(B))"
 | 
| 0 |     64 |   "SUM x:A. B"  => "Sum(A, %x. B)"
 | 
| 23 |     65 |   "A * B"       => "Sum(A, _K(B))"
 | 
| 0 |     66 | 
 | 
|  |     67 | rules
 | 
|  |     68 | 
 | 
|  |     69 |   (*Reduction: a weaker notion than equality;  a hack for simplification.
 | 
|  |     70 |     Reduce[a,b] means either that  a=b:A  for some A or else that "a" and "b"
 | 
|  |     71 |     are textually identical.*)
 | 
|  |     72 | 
 | 
|  |     73 |   (*does not verify a:A!  Sound because only trans_red uses a Reduce premise
 | 
|  |     74 |     No new theorems can be proved about the standard judgements.*)
 | 
|  |     75 |   refl_red "Reduce[a,a]"
 | 
|  |     76 |   red_if_equal "a = b : A ==> Reduce[a,b]"
 | 
|  |     77 |   trans_red "[| a = b : A;  Reduce[b,c] |] ==> a = c : A"
 | 
|  |     78 | 
 | 
|  |     79 |   (*Reflexivity*)
 | 
|  |     80 | 
 | 
|  |     81 |   refl_type "A type ==> A = A"
 | 
|  |     82 |   refl_elem "a : A ==> a = a : A"
 | 
|  |     83 | 
 | 
|  |     84 |   (*Symmetry*)
 | 
|  |     85 | 
 | 
|  |     86 |   sym_type  "A = B ==> B = A"
 | 
|  |     87 |   sym_elem  "a = b : A ==> b = a : A"
 | 
|  |     88 | 
 | 
|  |     89 |   (*Transitivity*)
 | 
|  |     90 | 
 | 
|  |     91 |   trans_type   "[| A = B;  B = C |] ==> A = C"
 | 
|  |     92 |   trans_elem   "[| a = b : A;  b = c : A |] ==> a = c : A"
 | 
|  |     93 | 
 | 
|  |     94 |   equal_types  "[| a : A;  A = B |] ==> a : B"
 | 
|  |     95 |   equal_typesL "[| a = b : A;  A = B |] ==> a = b : B"
 | 
|  |     96 | 
 | 
|  |     97 |   (*Substitution*)
 | 
|  |     98 | 
 | 
|  |     99 |   subst_type   "[| a : A;  !!z. z:A ==> B(z) type |] ==> B(a) type"
 | 
|  |    100 |   subst_typeL  "[| a = c : A;  !!z. z:A ==> B(z) = D(z) |] ==> B(a) = D(c)"
 | 
|  |    101 | 
 | 
|  |    102 |   subst_elem   "[| a : A;  !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)"
 | 
|  |    103 |   subst_elemL
 | 
|  |    104 |     "[| a=c : A;  !!z. z:A ==> b(z)=d(z) : B(z) |] ==> b(a)=d(c) : B(a)"
 | 
|  |    105 | 
 | 
|  |    106 | 
 | 
|  |    107 |   (*The type N -- natural numbers*)
 | 
|  |    108 | 
 | 
|  |    109 |   NF "N type"
 | 
|  |    110 |   NI0 "0 : N"
 | 
|  |    111 |   NI_succ "a : N ==> succ(a) : N"
 | 
|  |    112 |   NI_succL  "a = b : N ==> succ(a) = succ(b) : N"
 | 
|  |    113 | 
 | 
|  |    114 |   NE
 | 
| 1149 |    115 |    "[| p: N;  a: C(0);  !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] 
 | 
|  |    116 |    ==> rec(p, a, %u v.b(u,v)) : C(p)"
 | 
| 0 |    117 | 
 | 
|  |    118 |   NEL
 | 
| 1149 |    119 |    "[| p = q : N;  a = c : C(0);  
 | 
|  |    120 |       !!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u)) |] 
 | 
|  |    121 |    ==> rec(p, a, %u v.b(u,v)) = rec(q,c,d) : C(p)"
 | 
| 0 |    122 | 
 | 
|  |    123 |   NC0
 | 
| 1149 |    124 |    "[| a: C(0);  !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] 
 | 
|  |    125 |    ==> rec(0, a, %u v.b(u,v)) = a : C(0)"
 | 
| 0 |    126 | 
 | 
|  |    127 |   NC_succ
 | 
| 1149 |    128 |    "[| p: N;  a: C(0);  
 | 
|  |    129 |        !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] ==>  
 | 
|  |    130 |    rec(succ(p), a, %u v.b(u,v)) = b(p, rec(p, a, %u v.b(u,v))) : C(succ(p))"
 | 
| 0 |    131 | 
 | 
|  |    132 |   (*The fourth Peano axiom.  See page 91 of Martin-Lof's book*)
 | 
|  |    133 |   zero_ne_succ
 | 
|  |    134 |     "[| a: N;  0 = succ(a) : N |] ==> 0: F"
 | 
|  |    135 | 
 | 
|  |    136 | 
 | 
|  |    137 |   (*The Product of a family of types*)
 | 
|  |    138 | 
 | 
|  |    139 |   ProdF  "[| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A.B(x) type"
 | 
|  |    140 | 
 | 
|  |    141 |   ProdFL
 | 
| 1149 |    142 |    "[| A = C;  !!x. x:A ==> B(x) = D(x) |] ==> 
 | 
|  |    143 |    PROD x:A.B(x) = PROD x:C.D(x)"
 | 
| 0 |    144 | 
 | 
|  |    145 |   ProdI
 | 
|  |    146 |    "[| A type;  !!x. x:A ==> b(x):B(x)|] ==> lam x.b(x) : PROD x:A.B(x)"
 | 
|  |    147 | 
 | 
|  |    148 |   ProdIL
 | 
| 1149 |    149 |    "[| A type;  !!x. x:A ==> b(x) = c(x) : B(x)|] ==> 
 | 
|  |    150 |    lam x.b(x) = lam x.c(x) : PROD x:A.B(x)"
 | 
| 0 |    151 | 
 | 
|  |    152 |   ProdE  "[| p : PROD x:A.B(x);  a : A |] ==> p`a : B(a)"
 | 
|  |    153 |   ProdEL "[| p=q: PROD x:A.B(x);  a=b : A |] ==> p`a = q`b : B(a)"
 | 
|  |    154 | 
 | 
|  |    155 |   ProdC
 | 
| 1149 |    156 |    "[| a : A;  !!x. x:A ==> b(x) : B(x)|] ==> 
 | 
|  |    157 |    (lam x.b(x)) ` a = b(a) : B(a)"
 | 
| 0 |    158 | 
 | 
|  |    159 |   ProdC2
 | 
|  |    160 |    "p : PROD x:A.B(x) ==> (lam x. p`x) = p : PROD x:A.B(x)"
 | 
|  |    161 | 
 | 
|  |    162 | 
 | 
|  |    163 |   (*The Sum of a family of types*)
 | 
|  |    164 | 
 | 
|  |    165 |   SumF  "[| A type;  !!x. x:A ==> B(x) type |] ==> SUM x:A.B(x) type"
 | 
|  |    166 |   SumFL
 | 
|  |    167 |     "[| A = C;  !!x. x:A ==> B(x) = D(x) |] ==> SUM x:A.B(x) = SUM x:C.D(x)"
 | 
|  |    168 | 
 | 
|  |    169 |   SumI  "[| a : A;  b : B(a) |] ==> <a,b> : SUM x:A.B(x)"
 | 
|  |    170 |   SumIL "[| a=c:A;  b=d:B(a) |] ==> <a,b> = <c,d> : SUM x:A.B(x)"
 | 
|  |    171 | 
 | 
|  |    172 |   SumE
 | 
| 1149 |    173 |     "[| p: SUM x:A.B(x);  !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |] 
 | 
|  |    174 |     ==> split(p, %x y.c(x,y)) : C(p)"
 | 
| 0 |    175 | 
 | 
|  |    176 |   SumEL
 | 
| 1149 |    177 |     "[| p=q : SUM x:A.B(x); 
 | 
|  |    178 |        !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C(<x,y>)|] 
 | 
|  |    179 |     ==> split(p, %x y.c(x,y)) = split(q, % x y.d(x,y)) : C(p)"
 | 
| 0 |    180 | 
 | 
|  |    181 |   SumC
 | 
| 1149 |    182 |     "[| a: A;  b: B(a);  !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |] 
 | 
|  |    183 |     ==> split(<a,b>, %x y.c(x,y)) = c(a,b) : C(<a,b>)"
 | 
| 0 |    184 | 
 | 
|  |    185 |   fst_def   "fst(a) == split(a, %x y.x)"
 | 
|  |    186 |   snd_def   "snd(a) == split(a, %x y.y)"
 | 
|  |    187 | 
 | 
|  |    188 | 
 | 
|  |    189 |   (*The sum of two types*)
 | 
|  |    190 | 
 | 
|  |    191 |   PlusF   "[| A type;  B type |] ==> A+B type"
 | 
|  |    192 |   PlusFL  "[| A = C;  B = D |] ==> A+B = C+D"
 | 
|  |    193 | 
 | 
|  |    194 |   PlusI_inl   "[| a : A;  B type |] ==> inl(a) : A+B"
 | 
|  |    195 |   PlusI_inlL "[| a = c : A;  B type |] ==> inl(a) = inl(c) : A+B"
 | 
|  |    196 | 
 | 
|  |    197 |   PlusI_inr   "[| A type;  b : B |] ==> inr(b) : A+B"
 | 
|  |    198 |   PlusI_inrL "[| A type;  b = d : B |] ==> inr(b) = inr(d) : A+B"
 | 
|  |    199 | 
 | 
|  |    200 |   PlusE
 | 
| 1149 |    201 |     "[| p: A+B;  !!x. x:A ==> c(x): C(inl(x));  
 | 
|  |    202 |                 !!y. y:B ==> d(y): C(inr(y)) |] 
 | 
|  |    203 |     ==> when(p, %x.c(x), %y.d(y)) : C(p)"
 | 
| 0 |    204 | 
 | 
|  |    205 |   PlusEL
 | 
| 1149 |    206 |     "[| p = q : A+B;  !!x. x: A ==> c(x) = e(x) : C(inl(x));   
 | 
|  |    207 |                      !!y. y: B ==> d(y) = f(y) : C(inr(y)) |] 
 | 
|  |    208 |     ==> when(p, %x.c(x), %y.d(y)) = when(q, %x.e(x), %y.f(y)) : C(p)"
 | 
| 0 |    209 | 
 | 
|  |    210 |   PlusC_inl
 | 
| 1149 |    211 |     "[| a: A;  !!x. x:A ==> c(x): C(inl(x));  
 | 
|  |    212 |               !!y. y:B ==> d(y): C(inr(y)) |] 
 | 
|  |    213 |     ==> when(inl(a), %x.c(x), %y.d(y)) = c(a) : C(inl(a))"
 | 
| 0 |    214 | 
 | 
|  |    215 |   PlusC_inr
 | 
| 1149 |    216 |     "[| b: B;  !!x. x:A ==> c(x): C(inl(x));  
 | 
|  |    217 |               !!y. y:B ==> d(y): C(inr(y)) |] 
 | 
|  |    218 |     ==> when(inr(b), %x.c(x), %y.d(y)) = d(b) : C(inr(b))"
 | 
| 0 |    219 | 
 | 
|  |    220 | 
 | 
|  |    221 |   (*The type Eq*)
 | 
|  |    222 | 
 | 
|  |    223 |   EqF    "[| A type;  a : A;  b : A |] ==> Eq(A,a,b) type"
 | 
|  |    224 |   EqFL "[| A=B;  a=c: A;  b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)"
 | 
|  |    225 |   EqI "a = b : A ==> eq : Eq(A,a,b)"
 | 
|  |    226 |   EqE "p : Eq(A,a,b) ==> a = b : A"
 | 
|  |    227 | 
 | 
|  |    228 |   (*By equality of types, can prove C(p) from C(eq), an elimination rule*)
 | 
|  |    229 |   EqC "p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)"
 | 
|  |    230 | 
 | 
|  |    231 |   (*The type F*)
 | 
|  |    232 | 
 | 
|  |    233 |   FF "F type"
 | 
|  |    234 |   FE "[| p: F;  C type |] ==> contr(p) : C"
 | 
|  |    235 |   FEL  "[| p = q : F;  C type |] ==> contr(p) = contr(q) : C"
 | 
|  |    236 | 
 | 
|  |    237 |   (*The type T
 | 
|  |    238 |      Martin-Lof's book (page 68) discusses elimination and computation.
 | 
|  |    239 |      Elimination can be derived by computation and equality of types,
 | 
|  |    240 |      but with an extra premise C(x) type x:T.
 | 
|  |    241 |      Also computation can be derived from elimination. *)
 | 
|  |    242 | 
 | 
|  |    243 |   TF "T type"
 | 
|  |    244 |   TI "tt : T"
 | 
|  |    245 |   TE "[| p : T;  c : C(tt) |] ==> c : C(p)"
 | 
|  |    246 |   TEL "[| p = q : T;  c = d : C(tt) |] ==> c = d : C(p)"
 | 
|  |    247 |   TC "p : T ==> p = tt : T"
 | 
|  |    248 | end
 | 
|  |    249 | 
 | 
|  |    250 | 
 | 
|  |    251 | ML
 | 
|  |    252 | 
 | 
|  |    253 | val print_translation =
 | 
|  |    254 |   [("Prod", dependent_tr' ("@PROD", "@-->")),
 | 
|  |    255 |    ("Sum", dependent_tr' ("@SUM", "@*"))];
 | 
|  |    256 | 
 |