src/CTT/CTT.thy
author wenzelm
Wed, 16 Oct 2024 16:20:35 +0200
changeset 81173 6002cb6bfb0a
parent 80917 2a77bc3b4eac
permissions -rw-r--r--
performance tuning: cache markup and extern operations;
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
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     6
theory CTT
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     7
imports Pure
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     8
begin
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     9
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
    10
section \<open>Constructive Type Theory: axiomatic basis\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
    11
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    12
ML_file \<open>~~/src/Provers/typedsimp.ML\<close>
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
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    20
  \<comment> \<open>Judgments\<close>
80917
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    21
  Type      :: "t \<Rightarrow> prop"          (\<open>(\<open>notation=\<open>postfix Type\<close>\<close>_ type)\<close> [10] 5)
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    22
  Eqtype    :: "[t,t]\<Rightarrow>prop"        (\<open>(\<open>notation=\<open>infix Eqtype\<close>\<close>_ =/ _)\<close> [10,10] 5)
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    23
  Elem      :: "[i, t]\<Rightarrow>prop"       (\<open>(\<open>notation=\<open>infix Elem\<close>\<close>_ /: _)\<close> [10,10] 5)
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    24
  Eqelem    :: "[i,i,t]\<Rightarrow>prop"      (\<open>(\<open>notation=\<open>mixfix Eqelem\<close>\<close>_ =/ _ :/ _)\<close> [10,10,10] 5)
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    25
  Reduce    :: "[i,i]\<Rightarrow>prop"        (\<open>Reduce[_,_]\<close>)
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    26
  \<comment> \<open>Types for truth values\<close>
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    27
  F         :: "t"
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    28
  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
    29
  contr     :: "i\<Rightarrow>i"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
  tt        :: "i"
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    31
  \<comment> \<open>Natural numbers\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
  N         :: "t"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    33
  Zero      :: "i"                  (\<open>0\<close>)
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    34
  succ      :: "i\<Rightarrow>i"
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    35
  rec       :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow> i"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    36
  \<comment> \<open>Binary sum\<close>
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    37
  Plus      :: "[t,t]\<Rightarrow>t"           (infixr \<open>+\<close> 40)
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    38
  inl       :: "i\<Rightarrow>i"
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    39
  inr       :: "i\<Rightarrow>i"
60555
51a6997b1384 support 'when' statement, which corresponds to 'presume';
wenzelm
parents: 59780
diff changeset
    40
  "when"    :: "[i, i\<Rightarrow>i, i\<Rightarrow>i]\<Rightarrow>i"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    41
  \<comment> \<open>General sum and binary product\<close>
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    42
  Sum       :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
80917
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    43
  pair      :: "[i,i]\<Rightarrow>i"           (\<open>(\<open>indent=1 notation=\<open>mixfix pair\<close>\<close><_,/_>)\<close>)
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    44
  fst       :: "i\<Rightarrow>i"
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    45
  snd       :: "i\<Rightarrow>i"
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    46
  split     :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    47
  \<comment> \<open>General product and function space\<close>
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    48
  Prod      :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    49
  lambda    :: "(i \<Rightarrow> i) \<Rightarrow> i"      (binder \<open>\<^bold>\<lambda>\<close> 10)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    50
  app       :: "[i,i]\<Rightarrow>i"           (infixl \<open>`\<close> 60)
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    51
  \<comment> \<open>Equality type\<close>
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    52
  Eq        :: "[t,i,i]\<Rightarrow>t"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
  eq        :: "i"
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 14565
diff changeset
    54
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    55
text \<open>Some inexplicable syntactic dependencies; in particular, "0"
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    56
 must be introduced after the judgment forms.\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 14565
diff changeset
    58
syntax
80917
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    59
  "_PROD"   :: "[idt,t,t]\<Rightarrow>t"       (\<open>(\<open>indent=3 notation=\<open>binder \<Prod>\<close>\<close>\<Prod>_:_./ _)\<close> 10)
2a77bc3b4eac more inner syntax markup: minor object-logics;
wenzelm
parents: 80914
diff changeset
    60
  "_SUM"    :: "[idt,t,t]\<Rightarrow>t"       (\<open>(\<open>indent=3 notation=\<open>binder \<Sum>\<close>\<close>\<Sum>_:_./ _)\<close> 10)
80761
bc936d3d8b45 tuned, following be8c0e039a5e;
wenzelm
parents: 80754
diff changeset
    61
syntax_consts
bc936d3d8b45 tuned, following be8c0e039a5e;
wenzelm
parents: 80754
diff changeset
    62
  "_PROD" \<rightleftharpoons> Prod and
bc936d3d8b45 tuned, following be8c0e039a5e;
wenzelm
parents: 80754
diff changeset
    63
  "_SUM" \<rightleftharpoons> Sum
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
translations
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
    65
  "\<Prod>x:A. B" \<rightleftharpoons> "CONST Prod(A, \<lambda>x. B)"
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
    66
  "\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum(A, \<lambda>x. B)"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
    67
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    68
abbreviation Arrow :: "[t,t]\<Rightarrow>t"  (infixr \<open>\<longrightarrow>\<close> 30)
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    69
  where "A \<longrightarrow> B \<equiv> \<Prod>_:A. B"
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    70
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
    71
abbreviation Times :: "[t,t]\<Rightarrow>t"  (infixr \<open>\<times>\<close> 50)
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    72
  where "A \<times> B \<equiv> \<Sum>_:A. B"
10467
e6e7205e9e91 x-symbol support for Pi, Sigma, -->, : (membership)
paulson
parents: 3837
diff changeset
    73
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    74
text \<open>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    75
  Reduction: a weaker notion than equality;  a hack for simplification.
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    76
  \<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
    77
    that \<open>a\<close> and \<open>b\<close> are textually identical.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    79
  Does not verify \<open>a:A\<close>!  Sound because only \<open>trans_red\<close> uses a \<open>Reduce\<close>
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
    80
  premise. No new theorems can be proved about the standard judgments.
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    81
\<close>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    82
axiomatization
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    83
where
51308
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
    84
  refl_red: "\<And>a. Reduce[a,a]" and
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    85
  red_if_equal: "\<And>a b A. a = b : A \<Longrightarrow> Reduce[a,b]" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    86
  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
    87
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    88
  \<comment> \<open>Reflexivity\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    90
  refl_type: "\<And>A. A type \<Longrightarrow> A = A" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    91
  refl_elem: "\<And>a A. a : A \<Longrightarrow> a = a : A" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    93
  \<comment> \<open>Symmetry\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    95
  sym_type:  "\<And>A B. A = B \<Longrightarrow> B = A" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
    96
  sym_elem:  "\<And>a b A. a = b : A \<Longrightarrow> b = a : A" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
    98
  \<comment> \<open>Transitivity\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   100
  trans_type:   "\<And>A B C. \<lbrakk>A = B; B = C\<rbrakk> \<Longrightarrow> A = C" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   101
  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
   102
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   103
  equal_types:  "\<And>a A B. \<lbrakk>a : A; A = B\<rbrakk> \<Longrightarrow> a : B" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   104
  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
   105
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   106
  \<comment> \<open>Substitution\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   108
  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
   109
  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
   110
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   111
  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
   112
  subst_elemL:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   113
    "\<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
   114
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   116
  \<comment> \<open>The type \<open>N\<close> -- natural numbers\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
51308
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
   118
  NF: "N type" and
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
   119
  NI0: "0 : N" and
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   120
  NI_succ: "\<And>a. a : N \<Longrightarrow> succ(a) : N" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   121
  NI_succL:  "\<And>a b. a = b : N \<Longrightarrow> succ(a) = succ(b) : N" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   123
  NE:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   124
   "\<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
   125
   \<Longrightarrow> rec(p, a, \<lambda>u v. b(u,v)) : C(p)" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   126
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   127
  NEL:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   128
   "\<And>p q a b c d C. \<lbrakk>p = q : N; a = c : C(0);
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   129
      \<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
   130
   \<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
   131
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   132
  NC0:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   133
   "\<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
   134
   \<Longrightarrow> rec(0, a, \<lambda>u v. b(u,v)) = a : C(0)" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   136
  NC_succ:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   137
   "\<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
   138
   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
   139
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   140
  \<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
   141
  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
   142
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   144
  \<comment> \<open>The Product of a family of types\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   146
  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
   147
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   148
  ProdFL:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   149
    "\<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
   150
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   151
  ProdI:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   152
    "\<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
   153
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   154
  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
   155
    \<^bold>\<lambda>x. b(x) = \<^bold>\<lambda>x. c(x) : \<Prod>x:A. B(x)" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   157
  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
   158
  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
   159
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   160
  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
   161
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   162
  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
   163
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   165
  \<comment> \<open>The Sum of a family of types\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   167
  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
   168
  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
   169
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   170
  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
   171
  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
   172
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   173
  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
   174
    \<Longrightarrow> split(p, \<lambda>x y. c(x,y)) : C(p)" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   176
  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
   177
      \<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
   178
    \<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
   179
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   180
  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
   181
    \<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
   182
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   183
  fst_def:   "\<And>a. fst(a) \<equiv> split(a, \<lambda>x y. x)" and
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   184
  snd_def:   "\<And>a. snd(a) \<equiv> split(a, \<lambda>x y. y)" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   187
  \<comment> \<open>The sum of two types\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   189
  PlusF: "\<And>A B. \<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> A+B type" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   190
  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
   191
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   192
  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
   193
  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
   194
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   195
  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
   196
  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
   197
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   198
  PlusE:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   199
    "\<And>p c d A B C. \<lbrakk>p: A+B;
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   200
      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   201
      \<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
   202
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   203
  PlusEL:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   204
    "\<And>p q c d e f A B C. \<lbrakk>p = q : A+B;
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   205
      \<And>x. x: A \<Longrightarrow> c(x) = e(x) : C(inl(x));
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   206
      \<And>y. y: B \<Longrightarrow> d(y) = f(y) : C(inr(y))\<rbrakk>
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   207
    \<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
   208
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   209
  PlusC_inl:
64980
7dc25cf5793e misc tuning;
wenzelm
parents: 63505
diff changeset
   210
    "\<And>a c d A B C. \<lbrakk>a: A;
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   211
      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   212
      \<And>y. y:B \<Longrightarrow> d(y): C(inr(y)) \<rbrakk>
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   213
    \<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
   214
17441
5b5feca0344a converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   215
  PlusC_inr:
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   216
    "\<And>b c d A B C. \<lbrakk>b: B;
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   217
      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   218
      \<And>y. y:B \<Longrightarrow> d(y): C(inr(y))\<rbrakk>
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   219
    \<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
   220
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   222
  \<comment> \<open>The type \<open>Eq\<close>\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   223
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   224
  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
   225
  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
   226
  EqI: "\<And>a b A. a = b : A \<Longrightarrow> eq : Eq(A,a,b)" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   227
  EqE: "\<And>p a b A. p : Eq(A,a,b) \<Longrightarrow> a = b : A" and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   229
  \<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
   230
  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
   231
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   232
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   233
  \<comment> \<open>The type \<open>F\<close>\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   234
51308
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
   235
  FF: "F type" and
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   236
  FE: "\<And>p C. \<lbrakk>p: F; C type\<rbrakk> \<Longrightarrow> contr(p) : C" and
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   237
  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
   238
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   239
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   240
  \<comment> \<open>The type T\<close>
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67405
diff changeset
   241
  \<comment> \<open>Martin-Löf's book (page 68) discusses elimination and computation.
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   242
    Elimination can be derived by computation and equality of types,
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   243
    but with an extra premise \<open>C(x)\<close> type \<open>x:T\<close>.
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67405
diff changeset
   244
    Also computation can be derived from elimination.\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
51308
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
   246
  TF: "T type" and
51e158e988a5 eliminated legacy 'axioms';
wenzelm
parents: 48891
diff changeset
   247
  TI: "tt : T" and
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   248
  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
   249
  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
   250
  TC: "\<And>p. p : T \<Longrightarrow> p = tt : T"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   251
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   252
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   253
subsection "Tactics and derived rules for Constructive Type Theory"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   254
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   255
text \<open>Formation rules.\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   256
lemmas form_rls = NF ProdF SumF PlusF EqF FF TF
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   257
  and formL_rls = ProdFL SumFL PlusFL EqFL
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   258
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   259
text \<open>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   260
  Introduction rules. OMITTED:
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   261
  \<^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
   262
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   263
lemmas intr_rls = NI0 NI_succ ProdI SumI PlusI_inl PlusI_inr TI
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   264
  and intrL_rls = NI_succL ProdIL SumIL PlusI_inlL PlusI_inrL
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   265
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   266
text \<open>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   267
  Elimination rules. OMITTED:
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   268
  \<^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
   269
  \<^item> \<open>TE\<close>, because it does not involve a constructor.
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   270
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   271
lemmas elim_rls = NE ProdE SumE PlusE FE
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   272
  and elimL_rls = NEL ProdEL SumEL PlusEL FEL
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   273
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   274
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
   275
lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   276
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   277
text \<open>Rules with conclusion \<open>a:A\<close>, an elem judgment.\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   278
lemmas element_rls = intr_rls elim_rls
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>Definitions are (meta)equality axioms.\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   281
lemmas basic_defs = fst_def snd_def
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>Compare with standard version: \<open>B\<close> is applied to UNSIMPLIFIED expression!\<close>
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   284
lemma SumIL2: "\<lbrakk>c = a : A; d = b : B(a)\<rbrakk> \<Longrightarrow> <c,d> = <a,b> : Sum(A,B)"
65338
2ffda850f844 tuned proof;
wenzelm
parents: 64980
diff changeset
   285
  by (rule sym_elem) (rule SumIL; rule sym_elem)
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   286
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   287
lemmas intrL2_rls = NI_succL ProdIL SumIL2 PlusI_inlL PlusI_inrL
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   288
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   289
text \<open>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   290
  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
   291
  A more natural form of product elimination.
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   292
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   293
lemma subst_prodE:
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   294
  assumes "p: Prod(A,B)"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   295
    and "a: A"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   296
    and "\<And>z. z: B(a) \<Longrightarrow> c(z): C(z)"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   297
  shows "c(p`a): C(p`a)"
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   298
  by (rule assms ProdE)+
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   299
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   300
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   301
subsection \<open>Tactics for type checking\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   302
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   303
ML \<open>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   304
local
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   305
74299
16e5870fe21e more antiquotations;
wenzelm
parents: 69605
diff changeset
   306
fun is_rigid_elem \<^Const_>\<open>Elem for a _\<close> = not(is_Var (head_of a))
16e5870fe21e more antiquotations;
wenzelm
parents: 69605
diff changeset
   307
  | is_rigid_elem \<^Const_>\<open>Eqelem for a _ _\<close> = not(is_Var (head_of a))
16e5870fe21e more antiquotations;
wenzelm
parents: 69605
diff changeset
   308
  | is_rigid_elem \<^Const_>\<open>Type for a\<close> = not(is_Var (head_of a))
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   309
  | is_rigid_elem _ = false
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   310
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   311
in
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   312
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   313
(*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
   314
fun test_assume_tac ctxt = SUBGOAL (fn (prem, i) =>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   315
  if is_rigid_elem (Logic.strip_assums_concl prem)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   316
  then assume_tac ctxt i else no_tac)
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   317
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   318
fun ASSUME ctxt tf i = test_assume_tac ctxt i  ORELSE  tf i
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   319
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   320
end
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   321
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   322
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   323
text \<open>
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   324
  For simplification: type formation and checking,
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   325
  but no equalities between terms.
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   326
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   327
lemmas routine_rls = form_rls formL_rls refl_type element_rls
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   328
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   329
ML \<open>
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   330
fun routine_tac rls ctxt prems =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   331
  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
   332
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   333
(*Solve all subgoals "A type" using formation rules. *)
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   334
val form_net = Tactic.build_net @{thms form_rls};
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   335
fun form_tac ctxt =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   336
  REPEAT_FIRST (ASSUME ctxt (filt_resolve_from_net_tac ctxt 1 form_net));
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   337
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   338
(*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
   339
fun typechk_tac ctxt thms =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   340
  let val tac =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   341
    filt_resolve_from_net_tac ctxt 3
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   342
      (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
   343
  in  REPEAT_FIRST (ASSUME ctxt tac)  end
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   344
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   345
(*Solve a:A (a flexible, A rigid) by introduction rules.
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   346
  Cannot use stringtrees (filt_resolve_tac) since
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   347
  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
   348
fun intr_tac ctxt thms =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   349
  let val tac =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   350
    filt_resolve_from_net_tac ctxt 1
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   351
      (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
   352
  in  REPEAT_FIRST (ASSUME ctxt tac)  end
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   353
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   354
(*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
   355
fun equal_tac ctxt thms =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   356
  REPEAT_FIRST
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   357
    (ASSUME ctxt
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   358
      (filt_resolve_from_net_tac ctxt 3
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   359
        (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
   360
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   361
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   362
method_setup form = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\<close>
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   363
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
   364
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
   365
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
   366
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   367
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   368
subsection \<open>Simplification\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   369
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   370
text \<open>To simplify the type in a goal.\<close>
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   371
lemma replace_type: "\<lbrakk>B = A; a : A\<rbrakk> \<Longrightarrow> a : B"
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   372
  apply (rule equal_types)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   373
   apply (rule_tac [2] sym_type)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   374
   apply assumption+
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   375
  done
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>Simplify the parameter of a unary type operator.\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   378
lemma subst_eqtyparg:
23467
d1b97708d5eb tuned proofs -- avoid implicit prems;
wenzelm
parents: 22808
diff changeset
   379
  assumes 1: "a=c : A"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   380
    and 2: "\<And>z. z:A \<Longrightarrow> B(z) type"
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   381
  shows "B(a) = B(c)"
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   382
  apply (rule subst_typeL)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   383
   apply (rule_tac [2] refl_type)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   384
   apply (rule 1)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   385
  apply (erule 2)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   386
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   387
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   388
text \<open>Simplification rules for Constructive Type Theory.\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   389
lemmas reduction_rls = comp_rls [THEN trans_elem]
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   390
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   391
ML \<open>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   392
(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification.
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   393
  Uses other intro rules to avoid changing flexible goals.*)
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   394
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
   395
fun eqintr_tac ctxt =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 58977
diff changeset
   396
  REPEAT_FIRST (ASSUME ctxt (filt_resolve_from_net_tac ctxt 1 eqintr_net))
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   397
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   398
(** Tactics that instantiate CTT-rules.
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   399
    Vars in the given terms will be incremented!
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   400
    The (rtac EqE i) lets them apply to equality judgments. **)
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   401
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   402
fun NE_tac ctxt sp i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60555
diff changeset
   403
  TRY (resolve_tac ctxt @{thms EqE} i) THEN
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   404
  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
   405
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   406
fun SumE_tac ctxt sp i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60555
diff changeset
   407
  TRY (resolve_tac ctxt @{thms EqE} i) THEN
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   408
  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
   409
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   410
fun PlusE_tac ctxt sp i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60555
diff changeset
   411
  TRY (resolve_tac ctxt @{thms EqE} i) THEN
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   412
  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
   413
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   414
(** Predicate logic reasoning, WITH THINNING!!  Procedures adapted from NJ. **)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   415
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   416
(*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
   417
fun add_mp_tac ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60555
diff changeset
   418
  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
   419
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61378
diff changeset
   420
(*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
   421
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
   422
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   423
(*"safe" when regarded as predicate calculus rules*)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   424
val safe_brls = sort (make_ord lessb)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   425
    [ (true, @{thm FE}), (true,asm_rl),
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   426
      (false, @{thm ProdI}), (true, @{thm SumE}), (true, @{thm PlusE}) ]
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   427
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   428
val unsafe_brls =
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26956
diff changeset
   429
    [ (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
   430
      (true, @{thm subst_prodE}) ]
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   431
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   432
(*0 subgoals vs 1 or more*)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   433
val (safe0_brls, safep_brls) =
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67399
diff changeset
   434
    List.partition (curry (op =) 0 o subgoals_of_brl) safe_brls
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   435
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   436
fun safestep_tac ctxt thms i =
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   437
    form_tac ctxt ORELSE
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   438
    resolve_tac ctxt thms i  ORELSE
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   439
    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
   440
    DETERM (biresolve_tac ctxt safep_brls i)
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   441
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58889
diff changeset
   442
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
   443
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   444
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
   445
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   446
(*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
   447
fun pc_tac ctxt thms = DEPTH_SOLVE_1 o (step_tac ctxt thms)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   448
\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   449
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   450
method_setup eqintr = \<open>Scan.succeed (SIMPLE_METHOD o eqintr_tac)\<close>
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   451
method_setup NE = \<open>
74563
042041c0ebeb clarified modules;
wenzelm
parents: 74299
diff changeset
   452
  Scan.lift Parse.embedded_inner_syntax >> (fn s => fn ctxt => SIMPLE_METHOD' (NE_tac ctxt s))
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   453
\<close>
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   454
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
   455
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
   456
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   457
ML_file \<open>rew.ML\<close>
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   458
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
   459
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
   460
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   461
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 60754
diff changeset
   462
subsection \<open>The elimination rules for fst/snd\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   463
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   464
lemma SumE_fst: "p : Sum(A,B) \<Longrightarrow> fst(p) : A"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   465
  unfolding basic_defs
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   466
  apply (erule SumE)
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   467
  apply assumption
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   468
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   469
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   470
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
   471
lemma SumE_snd:
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   472
  assumes major: "p: Sum(A,B)"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   473
    and "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58976
diff changeset
   474
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   475
  shows "snd(p) : B(fst(p))"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   476
  unfolding basic_defs
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   477
  apply (rule major [THEN SumE])
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   478
  apply (rule SumC [THEN subst_eqtyparg, THEN replace_type])
63505
42e1dece537a misc tuning and modernization;
wenzelm
parents: 63120
diff changeset
   479
      apply (typechk assms)
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   480
  done
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   481
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   482
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   483
section \<open>The two-element type (booleans and conditionals)\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   484
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   485
definition Bool :: "t"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   486
  where "Bool \<equiv> T+T"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   487
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   488
definition true :: "i"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   489
  where "true \<equiv> inl(tt)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   490
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   491
definition false :: "i"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   492
  where "false \<equiv> inr(tt)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   493
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   494
definition cond :: "[i,i,i]\<Rightarrow>i"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   495
  where "cond(a,b,c) \<equiv> when(a, \<lambda>_. b, \<lambda>_. c)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   496
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   497
lemmas bool_defs = Bool_def true_def false_def cond_def
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   498
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   499
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   500
subsection \<open>Derivation of rules for the type \<open>Bool\<close>\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   501
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   502
text \<open>Formation rule.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   503
lemma boolF: "Bool type"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   504
  unfolding bool_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   505
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   506
text \<open>Introduction rules for \<open>true\<close>, \<open>false\<close>.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   507
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   508
lemma boolI_true: "true : Bool"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   509
  unfolding bool_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   510
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   511
lemma boolI_false: "false : Bool"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   512
  unfolding bool_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   513
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   514
text \<open>Elimination rule: typing of \<open>cond\<close>.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   515
lemma boolE: "\<lbrakk>p:Bool; a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(p,a,b) : C(p)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   516
  unfolding bool_defs
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   517
  apply (typechk; erule TE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   518
   apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   519
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   520
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   521
lemma boolEL: "\<lbrakk>p = q : Bool; a = c : C(true); b = d : C(false)\<rbrakk>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   522
  \<Longrightarrow> cond(p,a,b) = cond(q,c,d) : C(p)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   523
  unfolding bool_defs
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   524
  apply (rule PlusEL)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   525
    apply (erule asm_rl refl_elem [THEN TEL])+
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   526
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   527
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   528
text \<open>Computation rules for \<open>true\<close>, \<open>false\<close>.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   529
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   530
lemma boolC_true: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(true,a,b) = a : C(true)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   531
  unfolding bool_defs
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   532
  apply (rule comp_rls)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   533
    apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   534
   apply (erule_tac [!] TE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   535
   apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   536
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   537
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   538
lemma boolC_false: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(false,a,b) = b : C(false)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   539
  unfolding bool_defs
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   540
  apply (rule comp_rls)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   541
    apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   542
   apply (erule_tac [!] TE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   543
   apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   544
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   545
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   546
section \<open>Elementary arithmetic\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   547
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   548
subsection \<open>Arithmetic operators and their definitions\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   549
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   550
definition add :: "[i,i]\<Rightarrow>i"   (infixr \<open>#+\<close> 65)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   551
  where "a#+b \<equiv> rec(a, b, \<lambda>u v. succ(v))"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   552
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   553
definition diff :: "[i,i]\<Rightarrow>i"   (infixr \<open>-\<close> 65)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   554
  where "a-b \<equiv> rec(b, a, \<lambda>u v. rec(v, 0, \<lambda>x y. x))"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   555
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   556
definition absdiff :: "[i,i]\<Rightarrow>i"   (infixr \<open>|-|\<close> 65)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   557
  where "a|-|b \<equiv> (a-b) #+ (b-a)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   558
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   559
definition mult :: "[i,i]\<Rightarrow>i"   (infixr \<open>#*\<close> 70)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   560
  where "a#*b \<equiv> rec(a, 0, \<lambda>u v. b #+ v)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   561
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   562
definition mod :: "[i,i]\<Rightarrow>i"   (infixr \<open>mod\<close> 70)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   563
  where "a mod b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(v) |-| b, 0, \<lambda>x y. succ(v)))"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   564
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80761
diff changeset
   565
definition div :: "[i,i]\<Rightarrow>i"   (infixr \<open>div\<close> 70)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   566
  where "a div b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(u) mod b, succ(v), \<lambda>x y. v))"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   567
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   568
lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   569
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   570
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   571
subsection \<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   572
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   573
subsubsection \<open>Addition\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   574
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   575
text \<open>Typing of \<open>add\<close>: short and long versions.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   576
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   577
lemma add_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #+ b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   578
  unfolding arith_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   579
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   580
lemma add_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a #+ b = c #+ d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   581
  unfolding arith_defs by equal
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   582
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   583
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   584
text \<open>Computation for \<open>add\<close>: 0 and successor cases.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   585
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   586
lemma addC0: "b:N \<Longrightarrow> 0 #+ b = b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   587
  unfolding arith_defs by rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   588
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   589
lemma addC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #+ b = succ(a #+ b) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   590
  unfolding arith_defs by rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   591
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   592
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   593
subsubsection \<open>Multiplication\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   594
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   595
text \<open>Typing of \<open>mult\<close>: short and long versions.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   596
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   597
lemma mult_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   598
  unfolding arith_defs by (typechk add_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   599
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   600
lemma mult_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a #* b = c #* d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   601
  unfolding arith_defs by (equal add_typingL)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   602
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   603
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   604
text \<open>Computation for \<open>mult\<close>: 0 and successor cases.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   605
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   606
lemma multC0: "b:N \<Longrightarrow> 0 #* b = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   607
  unfolding arith_defs by rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   608
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   609
lemma multC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #* b = b #+ (a #* b) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   610
  unfolding arith_defs by rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   611
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   612
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   613
subsubsection \<open>Difference\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   614
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   615
text \<open>Typing of difference.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   616
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   617
lemma diff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a - b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   618
  unfolding arith_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   619
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   620
lemma diff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a - b = c - d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   621
  unfolding arith_defs by equal
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   622
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   623
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   624
text \<open>Computation for difference: 0 and successor cases.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   625
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   626
lemma diffC0: "a:N \<Longrightarrow> a - 0 = a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   627
  unfolding arith_defs by rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   628
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   629
text \<open>Note: \<open>rec(a, 0, \<lambda>z w.z)\<close> is \<open>pred(a).\<close>\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   630
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   631
lemma diff_0_eq_0: "b:N \<Longrightarrow> 0 - b = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   632
  unfolding arith_defs
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   633
  by (NE b) hyp_rew
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   634
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   635
text \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   636
  Essential to simplify FIRST!!  (Else we get a critical pair)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   637
  \<open>succ(a) - succ(b)\<close> rewrites to \<open>pred(succ(a) - b)\<close>.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   638
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   639
lemma diff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) - succ(b) = a - b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   640
  unfolding arith_defs
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   641
  apply hyp_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   642
  apply (NE b)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   643
    apply hyp_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   644
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   645
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   646
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   647
subsection \<open>Simplification\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   648
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   649
lemmas arith_typing_rls = add_typing mult_typing diff_typing
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   650
  and arith_congr_rls = add_typingL mult_typingL diff_typingL
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   651
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   652
lemmas congr_rls = arith_congr_rls intrL2_rls elimL_rls
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   653
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   654
lemmas arithC_rls =
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   655
  addC0 addC_succ
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   656
  multC0 multC_succ
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   657
  diffC0 diff_0_eq_0 diff_succ_succ
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   658
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   659
ML \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   660
  structure Arith_simp = TSimpFun(
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   661
    val refl = @{thm refl_elem}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   662
    val sym = @{thm sym_elem}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   663
    val trans = @{thm trans_elem}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   664
    val refl_red = @{thm refl_red}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   665
    val trans_red = @{thm trans_red}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   666
    val red_if_equal = @{thm red_if_equal}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   667
    val default_rls = @{thms arithC_rls comp_rls}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   668
    val routine_tac = routine_tac @{thms arith_typing_rls routine_rls}
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   669
  )
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   670
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   671
  fun arith_rew_tac ctxt prems =
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   672
    make_rew_tac ctxt (Arith_simp.norm_tac ctxt (@{thms congr_rls}, prems))
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   673
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   674
  fun hyp_arith_rew_tac ctxt prems =
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   675
    make_rew_tac ctxt
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   676
      (Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, @{thms congr_rls}, prems))
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   677
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   678
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   679
method_setup arith_rew = \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   680
  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (arith_rew_tac ctxt ths))
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   681
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   682
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   683
method_setup hyp_arith_rew = \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   684
  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_arith_rew_tac ctxt ths))
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   685
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   686
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   687
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   688
subsection \<open>Addition\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   689
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   690
text \<open>Associative law for addition.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   691
lemma add_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #+ b) #+ c = a #+ (b #+ c) : N"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   692
  by (NE a) hyp_arith_rew
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   693
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   694
text \<open>Commutative law for addition.  Can be proved using three inductions.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   695
  Must simplify after first induction!  Orientation of rewrites is delicate.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   696
lemma add_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #+ b = b #+ a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   697
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   698
    apply hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   699
   apply (rule sym_elem)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   700
   prefer 2
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   701
   apply (NE b)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   702
     prefer 4
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   703
     apply (NE b)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   704
       apply hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   705
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   706
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   707
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   708
subsection \<open>Multiplication\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   709
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   710
text \<open>Right annihilation in product.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   711
lemma mult_0_right: "a:N \<Longrightarrow> a #* 0 = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   712
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   713
    apply hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   714
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   715
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   716
text \<open>Right successor law for multiplication.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   717
lemma mult_succ_right: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* succ(b) = a #+ (a #* b) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   718
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   719
    apply (hyp_arith_rew add_assoc [THEN sym_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   720
  apply (assumption | rule add_commute mult_typingL add_typingL intrL_rls refl_elem)+
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   721
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   722
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   723
text \<open>Commutative law for multiplication.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   724
lemma mult_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* b = b #* a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   725
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   726
    apply (hyp_arith_rew mult_0_right mult_succ_right)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   727
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   728
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   729
text \<open>Addition distributes over multiplication.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   730
lemma add_mult_distrib: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #+ b) #* c = (a #* c) #+ (b #* c) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   731
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   732
    apply (hyp_arith_rew add_assoc [THEN sym_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   733
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   734
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   735
text \<open>Associative law for multiplication.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   736
lemma mult_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #* b) #* c = a #* (b #* c) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   737
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   738
    apply (hyp_arith_rew add_mult_distrib)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   739
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   740
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   741
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   742
subsection \<open>Difference\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   743
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   744
text \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   745
  Difference on natural numbers, without negative numbers
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   746
  \<^item> \<open>a - b = 0\<close>  iff  \<open>a \<le> b\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   747
  \<^item> \<open>a - b = succ(c)\<close> iff \<open>a > b\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   748
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   749
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   750
lemma diff_self_eq_0: "a:N \<Longrightarrow> a - a = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   751
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   752
    apply hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   753
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   754
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   755
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   756
lemma add_0_right: "\<lbrakk>c : N; 0 : N; c : N\<rbrakk> \<Longrightarrow> c #+ 0 = c : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   757
  by (rule addC0 [THEN [3] add_commute [THEN trans_elem]])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   758
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   759
text \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   760
  Addition is the inverse of subtraction: if \<open>b \<le> x\<close> then \<open>b #+ (x - b) = x\<close>.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   761
  An example of induction over a quantified formula (a product).
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   762
  Uses rewriting with a quantified, implicative inductive hypothesis.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   763
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   764
schematic_goal add_diff_inverse_lemma:
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   765
  "b:N \<Longrightarrow> ?a : \<Prod>x:N. Eq(N, b-x, 0) \<longrightarrow> Eq(N, b #+ (x-b), x)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   766
  apply (NE b)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   767
    \<comment> \<open>strip one "universal quantifier" but not the "implication"\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   768
    apply (rule_tac [3] intr_rls)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   769
    \<comment> \<open>case analysis on \<open>x\<close> in \<open>succ(u) \<le> x \<longrightarrow> succ(u) #+ (x - succ(u)) = x\<close>\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   770
     prefer 4
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   771
     apply (NE x)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   772
       apply assumption
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   773
    \<comment> \<open>Prepare for simplification of types -- the antecedent \<open>succ(u) \<le> x\<close>\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   774
      apply (rule_tac [2] replace_type)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   775
       apply (rule_tac [1] replace_type)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   776
        apply arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   777
    \<comment> \<open>Solves first 0 goal, simplifies others.  Two sugbgoals remain.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   778
    Both follow by rewriting, (2) using quantified induction hyp.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   779
   apply intr \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   780
    apply (hyp_arith_rew add_0_right)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   781
  apply assumption
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   782
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   783
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   784
text \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   785
  Version of above with premise \<open>b - a = 0\<close> i.e. \<open>a \<ge> b\<close>.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   786
  Using @{thm ProdE} does not work -- for \<open>?B(?a)\<close> is ambiguous.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   787
  Instead, @{thm add_diff_inverse_lemma} states the desired induction scheme;
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   788
  the use of \<open>THEN\<close> below instantiates Vars in @{thm ProdE} automatically.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   789
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   790
lemma add_diff_inverse: "\<lbrakk>a:N; b:N; b - a = 0 : N\<rbrakk> \<Longrightarrow> b #+ (a-b) = a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   791
  apply (rule EqE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   792
  apply (rule add_diff_inverse_lemma [THEN ProdE, THEN ProdE])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   793
    apply (assumption | rule EqI)+
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   794
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   795
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   796
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   797
subsection \<open>Absolute difference\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   798
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   799
text \<open>Typing of absolute difference: short and long versions.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   800
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   801
lemma absdiff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a |-| b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   802
  unfolding arith_defs by typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   803
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   804
lemma absdiff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a |-| b = c |-| d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   805
  unfolding arith_defs by equal
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   806
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   807
lemma absdiff_self_eq_0: "a:N \<Longrightarrow> a |-| a = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   808
  unfolding absdiff_def by (arith_rew diff_self_eq_0)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   809
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   810
lemma absdiffC0: "a:N \<Longrightarrow> 0 |-| a = a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   811
  unfolding absdiff_def by hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   812
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   813
lemma absdiff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) |-| succ(b)  =  a |-| b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   814
  unfolding absdiff_def by hyp_arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   815
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   816
text \<open>Note how easy using commutative laws can be?  ...not always...\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   817
lemma absdiff_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a |-| b = b |-| a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   818
  unfolding absdiff_def
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   819
  by (rule add_commute) (typechk diff_typing)
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   820
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   821
text \<open>If \<open>a + b = 0\<close> then \<open>a = 0\<close>. Surprisingly tedious.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   822
schematic_goal add_eq0_lemma: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> ?c : Eq(N,a#+b,0) \<longrightarrow> Eq(N,a,0)"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   823
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   824
    apply (rule_tac [3] replace_type)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   825
     apply arith_rew
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   826
  apply intr  \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   827
   apply (rule_tac [2] zero_ne_succ [THEN FE])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   828
     apply (erule_tac [3] EqE [THEN sym_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   829
    apply (typechk add_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   830
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   831
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   832
text \<open>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   833
  Version of above with the premise \<open>a + b = 0\<close>.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   834
  Again, resolution instantiates variables in @{thm ProdE}.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   835
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   836
lemma add_eq0: "\<lbrakk>a:N; b:N; a #+ b = 0 : N\<rbrakk> \<Longrightarrow> a = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   837
  apply (rule EqE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   838
  apply (rule add_eq0_lemma [THEN ProdE])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   839
    apply (rule_tac [3] EqI)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   840
    apply typechk
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   841
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   842
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   843
text \<open>Here is a lemma to infer \<open>a - b = 0\<close> and \<open>b - a = 0\<close> from \<open>a |-| b = 0\<close>, below.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   844
schematic_goal absdiff_eq0_lem:
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   845
  "\<lbrakk>a:N; b:N; a |-| b = 0 : N\<rbrakk> \<Longrightarrow> ?a : Eq(N, a-b, 0) \<times> Eq(N, b-a, 0)"
76381
2931d8331cc5 Beautification of some declarations
paulson <lp15@cam.ac.uk>
parents: 74563
diff changeset
   846
  unfolding absdiff_def
65447
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   847
  apply intr
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   848
   apply eqintr
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   849
   apply (rule_tac [2] add_eq0)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   850
     apply (rule add_eq0)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   851
       apply (rule_tac [6] add_commute [THEN trans_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   852
         apply (typechk diff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   853
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   854
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   855
text \<open>If \<open>a |-| b = 0\<close> then \<open>a = b\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   856
  proof: \<open>a - b = 0\<close> and \<open>b - a = 0\<close>, so \<open>b = a + (b - a) = a + 0 = a\<close>.
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   857
\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   858
lemma absdiff_eq0: "\<lbrakk>a |-| b = 0 : N; a:N; b:N\<rbrakk> \<Longrightarrow> a = b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   859
  apply (rule EqE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   860
  apply (rule absdiff_eq0_lem [THEN SumE])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   861
     apply eqintr
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   862
  apply (rule add_diff_inverse [THEN sym_elem, THEN trans_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   863
     apply (erule_tac [3] EqE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   864
    apply (hyp_arith_rew add_0_right)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   865
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   866
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   867
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   868
subsection \<open>Remainder and Quotient\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   869
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   870
text \<open>Typing of remainder: short and long versions.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   871
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   872
lemma mod_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a mod b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   873
  unfolding mod_def by (typechk absdiff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   874
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   875
lemma mod_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a mod b = c mod d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   876
  unfolding mod_def by (equal absdiff_typingL)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   877
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   878
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   879
text \<open>Computation for \<open>mod\<close>: 0 and successor cases.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   880
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   881
lemma modC0: "b:N \<Longrightarrow> 0 mod b = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   882
  unfolding mod_def by (rew absdiff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   883
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   884
lemma modC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   885
  succ(a) mod b = rec(succ(a mod b) |-| b, 0, \<lambda>x y. succ(a mod b)) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   886
  unfolding mod_def by (rew absdiff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   887
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   888
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   889
text \<open>Typing of quotient: short and long versions.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   890
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   891
lemma div_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a div b : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   892
  unfolding div_def by (typechk absdiff_typing mod_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   893
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   894
lemma div_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a div b = c div d : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   895
  unfolding div_def by (equal absdiff_typingL mod_typingL)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   896
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   897
lemmas div_typing_rls = mod_typing div_typing absdiff_typing
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   898
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   899
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   900
text \<open>Computation for quotient: 0 and successor cases.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   901
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   902
lemma divC0: "b:N \<Longrightarrow> 0 div b = 0 : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   903
  unfolding div_def by (rew mod_typing absdiff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   904
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   905
lemma divC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   906
  succ(a) div b = rec(succ(a) mod b, succ(a div b), \<lambda>x y. a div b) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   907
  unfolding div_def by (rew mod_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   908
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   909
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   910
text \<open>Version of above with same condition as the \<open>mod\<close> one.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   911
lemma divC_succ2: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   912
  succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), \<lambda>x y. a div b) : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   913
  apply (rule divC_succ [THEN trans_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   914
    apply (rew div_typing_rls modC_succ)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   915
  apply (NE "succ (a mod b) |-|b")
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   916
    apply (rew mod_typing div_typing absdiff_typing)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   917
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   918
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   919
text \<open>For case analysis on whether a number is 0 or a successor.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   920
lemma iszero_decidable: "a:N \<Longrightarrow> rec(a, inl(eq), \<lambda>ka kb. inr(<ka, eq>)) :
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   921
  Eq(N,a,0) + (\<Sum>x:N. Eq(N,a, succ(x)))"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   922
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   923
    apply (rule_tac [3] PlusI_inr)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   924
     apply (rule_tac [2] PlusI_inl)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   925
      apply eqintr
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   926
     apply equal
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   927
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   928
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   929
text \<open>Main Result. Holds when \<open>b\<close> is 0 since \<open>a mod 0 = a\<close> and \<open>a div 0 = 0\<close>.\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   930
lemma mod_div_equality: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a mod b #+ (a div b) #* b = a : N"
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   931
  apply (NE a)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   932
    apply (arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   933
  apply (rule EqE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   934
    \<comment> \<open>case analysis on \<open>succ(u mod b) |-| b\<close>\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   935
  apply (rule_tac a1 = "succ (u mod b) |-| b" in iszero_decidable [THEN PlusE])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   936
    apply (erule_tac [3] SumE)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   937
    apply (hyp_arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   938
    \<comment> \<open>Replace one occurrence of \<open>b\<close> by \<open>succ(u mod b)\<close>. Clumsy!\<close>
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   939
  apply (rule add_typingL [THEN trans_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   940
    apply (erule EqE [THEN absdiff_eq0, THEN sym_elem])
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   941
     apply (rule_tac [3] refl_elem)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   942
     apply (hyp_arith_rew div_typing_rls)
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   943
  done
fae6051ec192 clarified main CTT.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 65338
diff changeset
   944
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents: 17782
diff changeset
   945
end