src/CTT/CTT.thy
author wenzelm
Mon, 19 Sep 2011 22:13:51 +0200
changeset 45872 d88f7fc62a40
parent 41774 54b4686704af
child 49906 c0eafbd55de3
permissions -rw-r--r--
double clicks switch to document node buffer;
     1 (*  Title:      CTT/CTT.thy
     2     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     3     Copyright   1993  University of Cambridge
     4 *)
     5 
     6 header {* Constructive Type Theory *}
     7 
     8 theory CTT
     9 imports Pure
    10 uses "~~/src/Provers/typedsimp.ML" ("rew.ML")
    11 begin
    12 
    13 setup Pure_Thy.old_appl_syntax_setup
    14 
    15 typedecl i
    16 typedecl t
    17 typedecl o
    18 
    19 consts
    20   (*Types*)
    21   F         :: "t"
    22   T         :: "t"          (*F is empty, T contains one element*)
    23   contr     :: "i=>i"
    24   tt        :: "i"
    25   (*Natural numbers*)
    26   N         :: "t"
    27   succ      :: "i=>i"
    28   rec       :: "[i, i, [i,i]=>i] => i"
    29   (*Unions*)
    30   inl       :: "i=>i"
    31   inr       :: "i=>i"
    32   when      :: "[i, i=>i, i=>i]=>i"
    33   (*General Sum and Binary Product*)
    34   Sum       :: "[t, i=>t]=>t"
    35   fst       :: "i=>i"
    36   snd       :: "i=>i"
    37   split     :: "[i, [i,i]=>i] =>i"
    38   (*General Product and Function Space*)
    39   Prod      :: "[t, i=>t]=>t"
    40   (*Types*)
    41   Plus      :: "[t,t]=>t"           (infixr "+" 40)
    42   (*Equality type*)
    43   Eq        :: "[t,i,i]=>t"
    44   eq        :: "i"
    45   (*Judgements*)
    46   Type      :: "t => prop"          ("(_ type)" [10] 5)
    47   Eqtype    :: "[t,t]=>prop"        ("(_ =/ _)" [10,10] 5)
    48   Elem      :: "[i, t]=>prop"       ("(_ /: _)" [10,10] 5)
    49   Eqelem    :: "[i,i,t]=>prop"      ("(_ =/ _ :/ _)" [10,10,10] 5)
    50   Reduce    :: "[i,i]=>prop"        ("Reduce[_,_]")
    51   (*Types*)
    52 
    53   (*Functions*)
    54   lambda    :: "(i => i) => i"      (binder "lam " 10)
    55   app       :: "[i,i]=>i"           (infixl "`" 60)
    56   (*Natural numbers*)
    57   Zero      :: "i"                  ("0")
    58   (*Pairing*)
    59   pair      :: "[i,i]=>i"           ("(1<_,/_>)")
    60 
    61 syntax
    62   "_PROD"   :: "[idt,t,t]=>t"       ("(3PROD _:_./ _)" 10)
    63   "_SUM"    :: "[idt,t,t]=>t"       ("(3SUM _:_./ _)" 10)
    64 translations
    65   "PROD x:A. B" == "CONST Prod(A, %x. B)"
    66   "SUM x:A. B"  == "CONST Sum(A, %x. B)"
    67 
    68 abbreviation
    69   Arrow     :: "[t,t]=>t"  (infixr "-->" 30) where
    70   "A --> B == PROD _:A. B"
    71 abbreviation
    72   Times     :: "[t,t]=>t"  (infixr "*" 50) where
    73   "A * B == SUM _:A. B"
    74 
    75 notation (xsymbols)
    76   lambda  (binder "\<lambda>\<lambda>" 10) and
    77   Elem  ("(_ /\<in> _)" [10,10] 5) and
    78   Eqelem  ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and
    79   Arrow  (infixr "\<longrightarrow>" 30) and
    80   Times  (infixr "\<times>" 50)
    81 
    82 notation (HTML output)
    83   lambda  (binder "\<lambda>\<lambda>" 10) and
    84   Elem  ("(_ /\<in> _)" [10,10] 5) and
    85   Eqelem  ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and
    86   Times  (infixr "\<times>" 50)
    87 
    88 syntax (xsymbols)
    89   "_PROD"   :: "[idt,t,t] => t"     ("(3\<Pi> _\<in>_./ _)"    10)
    90   "_SUM"    :: "[idt,t,t] => t"     ("(3\<Sigma> _\<in>_./ _)" 10)
    91 
    92 syntax (HTML output)
    93   "_PROD"   :: "[idt,t,t] => t"     ("(3\<Pi> _\<in>_./ _)"    10)
    94   "_SUM"    :: "[idt,t,t] => t"     ("(3\<Sigma> _\<in>_./ _)" 10)
    95 
    96 axioms
    97 
    98   (*Reduction: a weaker notion than equality;  a hack for simplification.
    99     Reduce[a,b] means either that  a=b:A  for some A or else that "a" and "b"
   100     are textually identical.*)
   101 
   102   (*does not verify a:A!  Sound because only trans_red uses a Reduce premise
   103     No new theorems can be proved about the standard judgements.*)
   104   refl_red: "Reduce[a,a]"
   105   red_if_equal: "a = b : A ==> Reduce[a,b]"
   106   trans_red: "[| a = b : A;  Reduce[b,c] |] ==> a = c : A"
   107 
   108   (*Reflexivity*)
   109 
   110   refl_type: "A type ==> A = A"
   111   refl_elem: "a : A ==> a = a : A"
   112 
   113   (*Symmetry*)
   114 
   115   sym_type:  "A = B ==> B = A"
   116   sym_elem:  "a = b : A ==> b = a : A"
   117 
   118   (*Transitivity*)
   119 
   120   trans_type:   "[| A = B;  B = C |] ==> A = C"
   121   trans_elem:   "[| a = b : A;  b = c : A |] ==> a = c : A"
   122 
   123   equal_types:  "[| a : A;  A = B |] ==> a : B"
   124   equal_typesL: "[| a = b : A;  A = B |] ==> a = b : B"
   125 
   126   (*Substitution*)
   127 
   128   subst_type:   "[| a : A;  !!z. z:A ==> B(z) type |] ==> B(a) type"
   129   subst_typeL:  "[| a = c : A;  !!z. z:A ==> B(z) = D(z) |] ==> B(a) = D(c)"
   130 
   131   subst_elem:   "[| a : A;  !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)"
   132   subst_elemL:
   133     "[| a=c : A;  !!z. z:A ==> b(z)=d(z) : B(z) |] ==> b(a)=d(c) : B(a)"
   134 
   135 
   136   (*The type N -- natural numbers*)
   137 
   138   NF: "N type"
   139   NI0: "0 : N"
   140   NI_succ: "a : N ==> succ(a) : N"
   141   NI_succL:  "a = b : N ==> succ(a) = succ(b) : N"
   142 
   143   NE:
   144    "[| p: N;  a: C(0);  !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |]
   145    ==> rec(p, a, %u v. b(u,v)) : C(p)"
   146 
   147   NEL:
   148    "[| p = q : N;  a = c : C(0);
   149       !!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u)) |]
   150    ==> rec(p, a, %u v. b(u,v)) = rec(q,c,d) : C(p)"
   151 
   152   NC0:
   153    "[| a: C(0);  !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |]
   154    ==> rec(0, a, %u v. b(u,v)) = a : C(0)"
   155 
   156   NC_succ:
   157    "[| p: N;  a: C(0);
   158        !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u)) |] ==>
   159    rec(succ(p), a, %u v. b(u,v)) = b(p, rec(p, a, %u v. b(u,v))) : C(succ(p))"
   160 
   161   (*The fourth Peano axiom.  See page 91 of Martin-Lof's book*)
   162   zero_ne_succ:
   163     "[| a: N;  0 = succ(a) : N |] ==> 0: F"
   164 
   165 
   166   (*The Product of a family of types*)
   167 
   168   ProdF:  "[| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A. B(x) type"
   169 
   170   ProdFL:
   171    "[| A = C;  !!x. x:A ==> B(x) = D(x) |] ==>
   172    PROD x:A. B(x) = PROD x:C. D(x)"
   173 
   174   ProdI:
   175    "[| A type;  !!x. x:A ==> b(x):B(x)|] ==> lam x. b(x) : PROD x:A. B(x)"
   176 
   177   ProdIL:
   178    "[| A type;  !!x. x:A ==> b(x) = c(x) : B(x)|] ==>
   179    lam x. b(x) = lam x. c(x) : PROD x:A. B(x)"
   180 
   181   ProdE:  "[| p : PROD x:A. B(x);  a : A |] ==> p`a : B(a)"
   182   ProdEL: "[| p=q: PROD x:A. B(x);  a=b : A |] ==> p`a = q`b : B(a)"
   183 
   184   ProdC:
   185    "[| a : A;  !!x. x:A ==> b(x) : B(x)|] ==>
   186    (lam x. b(x)) ` a = b(a) : B(a)"
   187 
   188   ProdC2:
   189    "p : PROD x:A. B(x) ==> (lam x. p`x) = p : PROD x:A. B(x)"
   190 
   191 
   192   (*The Sum of a family of types*)
   193 
   194   SumF:  "[| A type;  !!x. x:A ==> B(x) type |] ==> SUM x:A. B(x) type"
   195   SumFL:
   196     "[| A = C;  !!x. x:A ==> B(x) = D(x) |] ==> SUM x:A. B(x) = SUM x:C. D(x)"
   197 
   198   SumI:  "[| a : A;  b : B(a) |] ==> <a,b> : SUM x:A. B(x)"
   199   SumIL: "[| a=c:A;  b=d:B(a) |] ==> <a,b> = <c,d> : SUM x:A. B(x)"
   200 
   201   SumE:
   202     "[| p: SUM x:A. B(x);  !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |]
   203     ==> split(p, %x y. c(x,y)) : C(p)"
   204 
   205   SumEL:
   206     "[| p=q : SUM x:A. B(x);
   207        !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C(<x,y>)|]
   208     ==> split(p, %x y. c(x,y)) = split(q, % x y. d(x,y)) : C(p)"
   209 
   210   SumC:
   211     "[| a: A;  b: B(a);  !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>) |]
   212     ==> split(<a,b>, %x y. c(x,y)) = c(a,b) : C(<a,b>)"
   213 
   214   fst_def:   "fst(a) == split(a, %x y. x)"
   215   snd_def:   "snd(a) == split(a, %x y. y)"
   216 
   217 
   218   (*The sum of two types*)
   219 
   220   PlusF:   "[| A type;  B type |] ==> A+B type"
   221   PlusFL:  "[| A = C;  B = D |] ==> A+B = C+D"
   222 
   223   PlusI_inl:   "[| a : A;  B type |] ==> inl(a) : A+B"
   224   PlusI_inlL: "[| a = c : A;  B type |] ==> inl(a) = inl(c) : A+B"
   225 
   226   PlusI_inr:   "[| A type;  b : B |] ==> inr(b) : A+B"
   227   PlusI_inrL: "[| A type;  b = d : B |] ==> inr(b) = inr(d) : A+B"
   228 
   229   PlusE:
   230     "[| p: A+B;  !!x. x:A ==> c(x): C(inl(x));
   231                 !!y. y:B ==> d(y): C(inr(y)) |]
   232     ==> when(p, %x. c(x), %y. d(y)) : C(p)"
   233 
   234   PlusEL:
   235     "[| p = q : A+B;  !!x. x: A ==> c(x) = e(x) : C(inl(x));
   236                      !!y. y: B ==> d(y) = f(y) : C(inr(y)) |]
   237     ==> when(p, %x. c(x), %y. d(y)) = when(q, %x. e(x), %y. f(y)) : C(p)"
   238 
   239   PlusC_inl:
   240     "[| a: A;  !!x. x:A ==> c(x): C(inl(x));
   241               !!y. y:B ==> d(y): C(inr(y)) |]
   242     ==> when(inl(a), %x. c(x), %y. d(y)) = c(a) : C(inl(a))"
   243 
   244   PlusC_inr:
   245     "[| b: B;  !!x. x:A ==> c(x): C(inl(x));
   246               !!y. y:B ==> d(y): C(inr(y)) |]
   247     ==> when(inr(b), %x. c(x), %y. d(y)) = d(b) : C(inr(b))"
   248 
   249 
   250   (*The type Eq*)
   251 
   252   EqF:    "[| A type;  a : A;  b : A |] ==> Eq(A,a,b) type"
   253   EqFL: "[| A=B;  a=c: A;  b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)"
   254   EqI: "a = b : A ==> eq : Eq(A,a,b)"
   255   EqE: "p : Eq(A,a,b) ==> a = b : A"
   256 
   257   (*By equality of types, can prove C(p) from C(eq), an elimination rule*)
   258   EqC: "p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)"
   259 
   260   (*The type F*)
   261 
   262   FF: "F type"
   263   FE: "[| p: F;  C type |] ==> contr(p) : C"
   264   FEL:  "[| p = q : F;  C type |] ==> contr(p) = contr(q) : C"
   265 
   266   (*The type T
   267      Martin-Lof's book (page 68) discusses elimination and computation.
   268      Elimination can be derived by computation and equality of types,
   269      but with an extra premise C(x) type x:T.
   270      Also computation can be derived from elimination. *)
   271 
   272   TF: "T type"
   273   TI: "tt : T"
   274   TE: "[| p : T;  c : C(tt) |] ==> c : C(p)"
   275   TEL: "[| p = q : T;  c = d : C(tt) |] ==> c = d : C(p)"
   276   TC: "p : T ==> p = tt : T"
   277 
   278 
   279 subsection "Tactics and derived rules for Constructive Type Theory"
   280 
   281 (*Formation rules*)
   282 lemmas form_rls = NF ProdF SumF PlusF EqF FF TF
   283   and formL_rls = ProdFL SumFL PlusFL EqFL
   284 
   285 (*Introduction rules
   286   OMITTED: EqI, because its premise is an eqelem, not an elem*)
   287 lemmas intr_rls = NI0 NI_succ ProdI SumI PlusI_inl PlusI_inr TI
   288   and intrL_rls = NI_succL ProdIL SumIL PlusI_inlL PlusI_inrL
   289 
   290 (*Elimination rules
   291   OMITTED: EqE, because its conclusion is an eqelem,  not an elem
   292            TE, because it does not involve a constructor *)
   293 lemmas elim_rls = NE ProdE SumE PlusE FE
   294   and elimL_rls = NEL ProdEL SumEL PlusEL FEL
   295 
   296 (*OMITTED: eqC are TC because they make rewriting loop: p = un = un = ... *)
   297 lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr
   298 
   299 (*rules with conclusion a:A, an elem judgement*)
   300 lemmas element_rls = intr_rls elim_rls
   301 
   302 (*Definitions are (meta)equality axioms*)
   303 lemmas basic_defs = fst_def snd_def
   304 
   305 (*Compare with standard version: B is applied to UNSIMPLIFIED expression! *)
   306 lemma SumIL2: "[| c=a : A;  d=b : B(a) |] ==> <c,d> = <a,b> : Sum(A,B)"
   307 apply (rule sym_elem)
   308 apply (rule SumIL)
   309 apply (rule_tac [!] sym_elem)
   310 apply assumption+
   311 done
   312 
   313 lemmas intrL2_rls = NI_succL ProdIL SumIL2 PlusI_inlL PlusI_inrL
   314 
   315 (*Exploit p:Prod(A,B) to create the assumption z:B(a).
   316   A more natural form of product elimination. *)
   317 lemma subst_prodE:
   318   assumes "p: Prod(A,B)"
   319     and "a: A"
   320     and "!!z. z: B(a) ==> c(z): C(z)"
   321   shows "c(p`a): C(p`a)"
   322 apply (rule assms ProdE)+
   323 done
   324 
   325 
   326 subsection {* Tactics for type checking *}
   327 
   328 ML {*
   329 
   330 local
   331 
   332 fun is_rigid_elem (Const("CTT.Elem",_) $ a $ _) = not(is_Var (head_of a))
   333   | is_rigid_elem (Const("CTT.Eqelem",_) $ a $ _ $ _) = not(is_Var (head_of a))
   334   | is_rigid_elem (Const("CTT.Type",_) $ a) = not(is_Var (head_of a))
   335   | is_rigid_elem _ = false
   336 
   337 in
   338 
   339 (*Try solving a:A or a=b:A by assumption provided a is rigid!*)
   340 val test_assume_tac = SUBGOAL(fn (prem,i) =>
   341     if is_rigid_elem (Logic.strip_assums_concl prem)
   342     then  assume_tac i  else  no_tac)
   343 
   344 fun ASSUME tf i = test_assume_tac i  ORELSE  tf i
   345 
   346 end;
   347 
   348 *}
   349 
   350 (*For simplification: type formation and checking,
   351   but no equalities between terms*)
   352 lemmas routine_rls = form_rls formL_rls refl_type element_rls
   353 
   354 ML {*
   355 local
   356   val equal_rls = @{thms form_rls} @ @{thms element_rls} @ @{thms intrL_rls} @
   357     @{thms elimL_rls} @ @{thms refl_elem}
   358 in
   359 
   360 fun routine_tac rls prems = ASSUME (filt_resolve_tac (prems @ rls) 4);
   361 
   362 (*Solve all subgoals "A type" using formation rules. *)
   363 val form_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac @{thms form_rls} 1));
   364 
   365 (*Type checking: solve a:A (a rigid, A flexible) by intro and elim rules. *)
   366 fun typechk_tac thms =
   367   let val tac = filt_resolve_tac (thms @ @{thms form_rls} @ @{thms element_rls}) 3
   368   in  REPEAT_FIRST (ASSUME tac)  end
   369 
   370 (*Solve a:A (a flexible, A rigid) by introduction rules.
   371   Cannot use stringtrees (filt_resolve_tac) since
   372   goals like ?a:SUM(A,B) have a trivial head-string *)
   373 fun intr_tac thms =
   374   let val tac = filt_resolve_tac(thms @ @{thms form_rls} @ @{thms intr_rls}) 1
   375   in  REPEAT_FIRST (ASSUME tac)  end
   376 
   377 (*Equality proving: solve a=b:A (where a is rigid) by long rules. *)
   378 fun equal_tac thms =
   379   REPEAT_FIRST (ASSUME (filt_resolve_tac (thms @ equal_rls) 3))
   380 
   381 end
   382 
   383 *}
   384 
   385 
   386 subsection {* Simplification *}
   387 
   388 (*To simplify the type in a goal*)
   389 lemma replace_type: "[| B = A;  a : A |] ==> a : B"
   390 apply (rule equal_types)
   391 apply (rule_tac [2] sym_type)
   392 apply assumption+
   393 done
   394 
   395 (*Simplify the parameter of a unary type operator.*)
   396 lemma subst_eqtyparg:
   397   assumes 1: "a=c : A"
   398     and 2: "!!z. z:A ==> B(z) type"
   399   shows "B(a)=B(c)"
   400 apply (rule subst_typeL)
   401 apply (rule_tac [2] refl_type)
   402 apply (rule 1)
   403 apply (erule 2)
   404 done
   405 
   406 (*Simplification rules for Constructive Type Theory*)
   407 lemmas reduction_rls = comp_rls [THEN trans_elem]
   408 
   409 ML {*
   410 (*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification.
   411   Uses other intro rules to avoid changing flexible goals.*)
   412 val eqintr_tac = REPEAT_FIRST (ASSUME (filt_resolve_tac (@{thm EqI} :: @{thms intr_rls}) 1))
   413 
   414 (** Tactics that instantiate CTT-rules.
   415     Vars in the given terms will be incremented!
   416     The (rtac EqE i) lets them apply to equality judgements. **)
   417 
   418 fun NE_tac ctxt sp i =
   419   TRY (rtac @{thm EqE} i) THEN res_inst_tac ctxt [(("p", 0), sp)] @{thm NE} i
   420 
   421 fun SumE_tac ctxt sp i =
   422   TRY (rtac @{thm EqE} i) THEN res_inst_tac ctxt [(("p", 0), sp)] @{thm SumE} i
   423 
   424 fun PlusE_tac ctxt sp i =
   425   TRY (rtac @{thm EqE} i) THEN res_inst_tac ctxt [(("p", 0), sp)] @{thm PlusE} i
   426 
   427 (** Predicate logic reasoning, WITH THINNING!!  Procedures adapted from NJ. **)
   428 
   429 (*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *)
   430 fun add_mp_tac i =
   431     rtac @{thm subst_prodE} i  THEN  assume_tac i  THEN  assume_tac i
   432 
   433 (*Finds P-->Q and P in the assumptions, replaces implication by Q *)
   434 fun mp_tac i = etac @{thm subst_prodE} i  THEN  assume_tac i
   435 
   436 (*"safe" when regarded as predicate calculus rules*)
   437 val safe_brls = sort (make_ord lessb)
   438     [ (true, @{thm FE}), (true,asm_rl),
   439       (false, @{thm ProdI}), (true, @{thm SumE}), (true, @{thm PlusE}) ]
   440 
   441 val unsafe_brls =
   442     [ (false, @{thm PlusI_inl}), (false, @{thm PlusI_inr}), (false, @{thm SumI}),
   443       (true, @{thm subst_prodE}) ]
   444 
   445 (*0 subgoals vs 1 or more*)
   446 val (safe0_brls, safep_brls) =
   447     List.partition (curry (op =) 0 o subgoals_of_brl) safe_brls
   448 
   449 fun safestep_tac thms i =
   450     form_tac  ORELSE
   451     resolve_tac thms i  ORELSE
   452     biresolve_tac safe0_brls i  ORELSE  mp_tac i  ORELSE
   453     DETERM (biresolve_tac safep_brls i)
   454 
   455 fun safe_tac thms i = DEPTH_SOLVE_1 (safestep_tac thms i)
   456 
   457 fun step_tac thms = safestep_tac thms  ORELSE'  biresolve_tac unsafe_brls
   458 
   459 (*Fails unless it solves the goal!*)
   460 fun pc_tac thms = DEPTH_SOLVE_1 o (step_tac thms)
   461 *}
   462 
   463 use "rew.ML"
   464 
   465 
   466 subsection {* The elimination rules for fst/snd *}
   467 
   468 lemma SumE_fst: "p : Sum(A,B) ==> fst(p) : A"
   469 apply (unfold basic_defs)
   470 apply (erule SumE)
   471 apply assumption
   472 done
   473 
   474 (*The first premise must be p:Sum(A,B) !!*)
   475 lemma SumE_snd:
   476   assumes major: "p: Sum(A,B)"
   477     and "A type"
   478     and "!!x. x:A ==> B(x) type"
   479   shows "snd(p) : B(fst(p))"
   480   apply (unfold basic_defs)
   481   apply (rule major [THEN SumE])
   482   apply (rule SumC [THEN subst_eqtyparg, THEN replace_type])
   483   apply (tactic {* typechk_tac @{thms assms} *})
   484   done
   485 
   486 end