src/CTT/ex/Elimination.thy
author paulson <lp15@cam.ac.uk>
Fri, 08 Aug 2025 16:46:03 +0100
changeset 82969 dedd9d13c79c
parent 76539 8c94ca4dd035
permissions -rw-r--r--
Generalised a theorem about Lebesgue integration
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     1
(*  Title:      CTT/ex/Elimination.thy
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     3
    Copyright   1991  University of Cambridge
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     4
76063
24c9f56aa035 proper umlauts;
wenzelm
parents: 69593
diff changeset
     5
Some examples taken from P. Martin-Löf, Intuitionistic type theory
35762
af3ff2ba4c54 removed old CVS Ids;
wenzelm
parents: 19761
diff changeset
     6
(Bibliopolis, 1984).
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     7
*)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     8
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
     9
section \<open>Examples with elimination rules\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    10
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    11
theory Elimination
58974
cbc2ac19d783 simplifie sessions;
wenzelm
parents: 58972
diff changeset
    12
imports "../CTT"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    13
begin
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    14
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    15
text \<open>This finds the functions fst and snd!\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    16
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    17
schematic_goal [folded basic_defs]: "A type \<Longrightarrow> ?a : (A \<times> A) \<longrightarrow> A"
58972
5b026cfc5f04 more Isar proof methods;
wenzelm
parents: 58963
diff changeset
    18
apply pc
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    19
done
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    20
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    21
schematic_goal [folded basic_defs]: "A type \<Longrightarrow> ?a : (A \<times> A) \<longrightarrow> A"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    22
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    23
  back
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    24
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    25
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    26
text \<open>Double negation of the Excluded Middle\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    27
schematic_goal "A type \<Longrightarrow> ?a : ((A + (A\<longrightarrow>F)) \<longrightarrow> F) \<longrightarrow> F"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    28
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    29
  apply (rule ProdE)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    30
   apply assumption
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    31
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    32
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    33
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    34
text \<open>Experiment: the proof above in Isar\<close>
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    35
lemma
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    36
  assumes "A type" shows "(\<^bold>\<lambda>f. f ` inr(\<^bold>\<lambda>y. f ` inl(y))) : ((A + (A\<longrightarrow>F)) \<longrightarrow> F) \<longrightarrow> F"
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    37
proof intr
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    38
  fix f
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    39
  assume f: "f : A + (A \<longrightarrow> F) \<longrightarrow> F" 
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    40
  with assms have "inr(\<^bold>\<lambda>y. f ` inl(y)) : A + (A \<longrightarrow> F)"
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    41
    by pc
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    42
  then show "f ` inr(\<^bold>\<lambda>y. f ` inl(y)) : F" 
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    43
    by (rule ProdE [OF f])
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    44
qed (rule assms)+
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    45
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    46
schematic_goal "\<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> ?a : (A \<times> B) \<longrightarrow> (B \<times> A)"
58972
5b026cfc5f04 more Isar proof methods;
wenzelm
parents: 58963
diff changeset
    47
apply pc
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    48
done
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    49
(*The sequent version (ITT) could produce an interesting alternative
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    50
  by backtracking.  No longer.*)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    51
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    52
text \<open>Binary sums and products\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    53
schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A + B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> C) \<times> (B \<longrightarrow> C)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    54
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    55
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    56
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    57
(*A distributive law*)
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    58
schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : A \<times> (B + C) \<longrightarrow> (A \<times> B + A \<times> C)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    59
  by pc
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    60
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    61
(*more general version, same proof*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    62
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    63
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
    64
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
    65
    and "\<And>x. x:A \<Longrightarrow> C(x) type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    66
  shows "?a : (\<Sum>x:A. B(x) + C(x)) \<longrightarrow> (\<Sum>x:A. B(x)) + (\<Sum>x:A. C(x))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    67
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    68
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    69
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    70
text \<open>Construction of the currying functional\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    71
schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A \<times> B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> (B \<longrightarrow> C))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    72
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    73
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    74
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    75
(*more general goal with same proof*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    76
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    77
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
    78
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    79
    and "\<And>z. z: (\<Sum>x:A. B(x)) \<Longrightarrow> C(z) type"
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    80
  shows "?a : \<Prod>f: (\<Prod>z : (\<Sum>x:A . B(x)) . C(z)).
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    81
                      (\<Prod>x:A . \<Prod>y:B(x) . C(<x,y>))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    82
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    83
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    84
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
    85
text \<open>Martin-Löf (1984), page 48: axiom of sum-elimination (uncurry)\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    86
schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A \<longrightarrow> (B \<longrightarrow> C)) \<longrightarrow> (A \<times> B \<longrightarrow> C)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    87
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    88
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    89
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    90
(*more general goal with same proof*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
    91
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    92
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
    93
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    94
    and "\<And>z. z: (\<Sum>x:A . B(x)) \<Longrightarrow> C(z) type"
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    95
  shows "?a : (\<Prod>x:A . \<Prod>y:B(x) . C(<x,y>))
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    96
        \<longrightarrow> (\<Prod>z : (\<Sum>x:A . B(x)) . C(z))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    97
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    98
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    99
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   100
text \<open>Function application\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   101
schematic_goal "\<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> ?a : ((A \<longrightarrow> B) \<times> A) \<longrightarrow> B"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   102
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   103
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   104
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   105
text \<open>Basic test of quantifier reasoning\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   106
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   107
  assumes "A type"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   108
    and "B type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   109
    and "\<And>x y. \<lbrakk>x:A; y:B\<rbrakk> \<Longrightarrow> C(x,y) type"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   110
  shows
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   111
    "?a :     (\<Sum>y:B . \<Prod>x:A . C(x,y))
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   112
          \<longrightarrow> (\<Prod>x:A . \<Sum>y:B . C(x,y))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   113
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   114
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   115
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   116
text \<open>Martin-Löf (1984) pages 36-7: the combinator S\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   117
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   118
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   119
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   120
    and "\<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> C(x,y) type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   121
  shows "?a :    (\<Prod>x:A. \<Prod>y:B(x). C(x,y))
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   122
             \<longrightarrow> (\<Prod>f: (\<Prod>x:A. B(x)). \<Prod>x:A. C(x, f`x))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   123
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   124
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   125
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   126
text \<open>Martin-Löf (1984) page 58: the axiom of disjunction elimination\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   127
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   128
  assumes "A type"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   129
    and "B type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   130
    and "\<And>z. z: A+B \<Longrightarrow> C(z) type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   131
  shows "?a : (\<Prod>x:A. C(inl(x))) \<longrightarrow> (\<Prod>y:B. C(inr(y)))
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   132
          \<longrightarrow> (\<Prod>z: A+B. C(z))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   133
  apply (pc assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   134
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   135
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   136
(*towards AXIOM OF CHOICE*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   137
schematic_goal [folded basic_defs]:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   138
  "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A \<longrightarrow> B \<times> C) \<longrightarrow> (A \<longrightarrow> B) \<times> (A \<longrightarrow> C)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   139
  apply pc
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   140
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   141
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   142
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   143
(*Martin-Löf (1984) page 50*)
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   144
text \<open>AXIOM OF CHOICE!  Delicate use of elimination rules\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   145
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   146
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   147
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   148
    and "\<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> C(x,y) type"
64980
7dc25cf5793e misc tuning;
wenzelm
parents: 63505
diff changeset
   149
  shows "?a : (\<Prod>x:A. \<Sum>y:B(x). C(x,y)) \<longrightarrow> (\<Sum>f: (\<Prod>x:A. B(x)). \<Prod>x:A. C(x, f`x))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   150
  apply (intr assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   151
   prefer 2 apply add_mp
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   152
   prefer 2 apply add_mp
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   153
   apply (erule SumE_fst)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   154
  apply (rule replace_type)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   155
   apply (rule subst_eqtyparg)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   156
    apply (rule comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   157
     apply (rule_tac [4] SumE_snd)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   158
       apply (typechk SumE_fst assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   159
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   160
76520
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   161
text \<open>A structured proof of AC\<close>
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   162
lemma Axiom_of_Choice:
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   163
  assumes "A type"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   164
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   165
    and "\<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> C(x,y) type"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   166
  shows "(\<^bold>\<lambda>f. <\<^bold>\<lambda>x. fst(f`x), \<^bold>\<lambda>x. snd(f`x)>) 
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   167
        : (\<Prod>x:A. \<Sum>y:B(x). C(x,y)) \<longrightarrow> (\<Sum>f: (\<Prod>x:A. B(x)). \<Prod>x:A. C(x, f`x))"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   168
proof (intr assms)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   169
  fix f a
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   170
  assume f: "f : \<Prod>x:A. Sum(B(x), C(x))" and "a : A" 
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   171
  then have fa: "f`a : Sum(B(a), C(a))"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   172
    by (rule ProdE)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   173
  then show "fst(f ` a) : B(a)" 
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   174
    by (rule SumE_fst)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   175
  have "snd(f ` a) : C(a, fst(f ` a))"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   176
    by (rule SumE_snd [OF fa]) (typechk SumE_fst assms \<open>a : A\<close>)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   177
  moreover have "(\<^bold>\<lambda>x. fst(f ` x)) ` a = fst(f ` a) : B(a)"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   178
    by (rule ProdC [OF \<open>a : A\<close>]) (typechk SumE_fst f)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   179
  ultimately show "snd(f`a) : C(a, (\<^bold>\<lambda>x. fst(f ` x)) ` a)"
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   180
    by (intro replace_type [OF subst_eqtyparg]) (typechk SumE_fst assms \<open>a : A\<close>)
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   181
qed
4d6d8dfd2cd2 Added an example for Isabelle/CTT
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
   182
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   183
text \<open>Axiom of choice.  Proof without fst, snd.  Harder still!\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   184
schematic_goal [folded basic_defs]:
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   185
  assumes "A type"
58977
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   186
    and "\<And>x. x:A \<Longrightarrow> B(x) type"
9576b510f6a2 more symbols;
wenzelm
parents: 58974
diff changeset
   187
    and "\<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> C(x,y) type"
64980
7dc25cf5793e misc tuning;
wenzelm
parents: 63505
diff changeset
   188
  shows "?a : (\<Prod>x:A. \<Sum>y:B(x). C(x,y)) \<longrightarrow> (\<Sum>f: (\<Prod>x:A. B(x)). \<Prod>x:A. C(x, f`x))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   189
  apply (intr assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   190
    (*Must not use add_mp as subst_prodE hides the construction.*)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   191
   apply (rule ProdE [THEN SumE])
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   192
     apply assumption
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   193
    apply assumption
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   194
   apply assumption
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   195
  apply (rule replace_type)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   196
   apply (rule subst_eqtyparg)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   197
    apply (rule comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   198
     apply (erule_tac [4] ProdE [THEN SumE])
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   199
      apply (typechk assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   200
  apply (rule replace_type)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   201
   apply (rule subst_eqtyparg)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   202
    apply (rule comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   203
      apply (typechk assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   204
  apply assumption
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   205
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   206
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76520
diff changeset
   207
text \<open>Example of sequent-style deduction\<close>
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   208
  (*When splitting z:A \<times> B, the assumption C(z) is affected;  ?a becomes
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   209
    \<^bold>\<lambda>u. split(u,\<lambda>v w.split(v,\<lambda>x y.\<^bold> \<lambda>z. <x,<y,z>>) ` w)     *)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 60770
diff changeset
   210
schematic_goal
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   211
  assumes "A type"
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   212
    and "B type"
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   213
    and "\<And>z. z:A \<times> B \<Longrightarrow> C(z) type"
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
   214
  shows "?a : (\<Sum>z:A \<times> B. C(z)) \<longrightarrow> (\<Sum>u:A. \<Sum>v:B. C(<u,v>))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   215
  apply (rule intr_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   216
   apply (tactic \<open>biresolve_tac \<^context> safe_brls 2\<close>)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   217
    (*Now must convert assumption C(z) into antecedent C(<kd,ke>) *)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   218
   apply (rule_tac [2] a = "y" in ProdE)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   219
    apply (typechk assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   220
  apply (rule SumE, assumption)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   221
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   222
     defer 1
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   223
     apply assumption+
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   224
  apply (typechk assms)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   225
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   226
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   227
end