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