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