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