| author | wenzelm | 
| Tue, 26 Mar 2024 17:06:23 +0100 | |
| changeset 80008 | 914c4a81027d | 
| parent 78099 | 4d9349989d94 | 
| child 80754 | 701912f5645a | 
| permissions | -rw-r--r-- | 
| 41777 | 1 | (* Title: ZF/OrdQuant.thy | 
| 2469 | 2 | Authors: Krzysztof Grabczewski and L C Paulson | 
| 3 | *) | |
| 4 | ||
| 60770 | 5 | section \<open>Special quantifiers\<close> | 
| 13253 | 6 | |
| 16417 | 7 | theory OrdQuant imports Ordinal begin | 
| 2469 | 8 | |
| 60770 | 9 | subsection \<open>Quantifiers and union operator for ordinals\<close> | 
| 13253 | 10 | |
| 24893 | 11 | definition | 
| 2469 | 12 | (* Ordinal Quantifiers *) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 13 | oall :: "[i, i \<Rightarrow> o] \<Rightarrow> o" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 14 | "oall(A, P) \<equiv> \<forall>x. x<A \<longrightarrow> P(x)" | 
| 13298 | 15 | |
| 24893 | 16 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 17 | oex :: "[i, i \<Rightarrow> o] \<Rightarrow> o" where | 
| 76214 | 18 | "oex(A, P) \<equiv> \<exists>x. x<A \<and> P(x)" | 
| 2469 | 19 | |
| 24893 | 20 | definition | 
| 2469 | 21 | (* Ordinal Union *) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 22 | OUnion :: "[i, i \<Rightarrow> i] \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 23 |     "OUnion(i,B) \<equiv> {z: \<Union>x\<in>i. B(x). Ord(i)}"
 | 
| 13298 | 24 | |
| 2469 | 25 | syntax | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 26 | "_oall" :: "[idt, i, o] \<Rightarrow> o" (\<open>(3\<forall>_<_./ _)\<close> 10) | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 27 | "_oex" :: "[idt, i, o] \<Rightarrow> o" (\<open>(3\<exists>_<_./ _)\<close> 10) | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 28 | "_OUNION" :: "[idt, i, i] \<Rightarrow> i" (\<open>(3\<Union>_<_./ _)\<close> 10) | 
| 61396 | 29 | translations | 
| 30 | "\<forall>x<a. P" \<rightleftharpoons> "CONST oall(a, \<lambda>x. P)" | |
| 31 | "\<exists>x<a. P" \<rightleftharpoons> "CONST oex(a, \<lambda>x. P)" | |
| 32 | "\<Union>x<a. B" \<rightleftharpoons> "CONST OUnion(a, \<lambda>x. B)" | |
| 12620 | 33 | |
| 34 | ||
| 60770 | 35 | subsubsection \<open>simplification of the new quantifiers\<close> | 
| 12825 | 36 | |
| 37 | ||
| 13169 | 38 | (*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize | 
| 13298 | 39 | is proved. Ord_atomize would convert this rule to | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 40 | x < 0 \<Longrightarrow> P(x) \<equiv> True, which causes dire effects!*) | 
| 46820 | 41 | lemma [simp]: "(\<forall>x<0. P(x))" | 
| 13298 | 42 | by (simp add: oall_def) | 
| 12825 | 43 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 44 | lemma [simp]: "\<not>(\<exists>x<0. P(x))" | 
| 13298 | 45 | by (simp add: oex_def) | 
| 12825 | 46 | |
| 76214 | 47 | lemma [simp]: "(\<forall>x<succ(i). P(x)) <-> (Ord(i) \<longrightarrow> P(i) \<and> (\<forall>x<i. P(x)))" | 
| 13298 | 48 | apply (simp add: oall_def le_iff) | 
| 49 | apply (blast intro: lt_Ord2) | |
| 12825 | 50 | done | 
| 51 | ||
| 76214 | 52 | lemma [simp]: "(\<exists>x<succ(i). P(x)) <-> (Ord(i) \<and> (P(i) | (\<exists>x<i. P(x))))" | 
| 13298 | 53 | apply (simp add: oex_def le_iff) | 
| 54 | apply (blast intro: lt_Ord2) | |
| 12825 | 55 | done | 
| 56 | ||
| 60770 | 57 | subsubsection \<open>Union over ordinals\<close> | 
| 13118 | 58 | |
| 12620 | 59 | lemma Ord_OUN [intro,simp]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 60 | "\<lbrakk>\<And>x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))" | 
| 13298 | 61 | by (simp add: OUnion_def ltI Ord_UN) | 
| 12620 | 62 | |
| 63 | lemma OUN_upper_lt: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 64 | "\<lbrakk>a<A; i < b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))" | 
| 12620 | 65 | by (unfold OUnion_def lt_def, blast ) | 
| 66 | ||
| 67 | lemma OUN_upper_le: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 68 | "\<lbrakk>a<A; i\<le>b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))" | 
| 12820 | 69 | apply (unfold OUnion_def, auto) | 
| 12620 | 70 | apply (rule UN_upper_le ) | 
| 13298 | 71 | apply (auto simp add: lt_def) | 
| 12620 | 72 | done | 
| 2469 | 73 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 74 | lemma Limit_OUN_eq: "Limit(i) \<Longrightarrow> (\<Union>x<i. x) = i" | 
| 12620 | 75 | by (simp add: OUnion_def Limit_Union_eq Limit_is_Ord) | 
| 76 | ||
| 46820 | 77 | (* No < version of this theorem: consider that @{term"(\<Union>i\<in>nat.i)=nat"}! *)
 | 
| 12620 | 78 | lemma OUN_least: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 79 | "(\<And>x. x<A \<Longrightarrow> B(x) \<subseteq> C) \<Longrightarrow> (\<Union>x<A. B(x)) \<subseteq> C" | 
| 12620 | 80 | by (simp add: OUnion_def UN_least ltI) | 
| 81 | ||
| 82 | lemma OUN_least_le: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 83 | "\<lbrakk>Ord(i); \<And>x. x<A \<Longrightarrow> b(x) \<le> i\<rbrakk> \<Longrightarrow> (\<Union>x<A. b(x)) \<le> i" | 
| 12620 | 84 | by (simp add: OUnion_def UN_least_le ltI Ord_0_le) | 
| 85 | ||
| 86 | lemma le_implies_OUN_le_OUN: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 87 | "\<lbrakk>\<And>x. x<A \<Longrightarrow> c(x) \<le> d(x)\<rbrakk> \<Longrightarrow> (\<Union>x<A. c(x)) \<le> (\<Union>x<A. d(x))" | 
| 12620 | 88 | by (blast intro: OUN_least_le OUN_upper_le le_Ord2 Ord_OUN) | 
| 89 | ||
| 90 | lemma OUN_UN_eq: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 91 | "(\<And>x. x \<in> A \<Longrightarrow> Ord(B(x))) | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 92 | \<Longrightarrow> (\<Union>z < (\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z < B(x). C(z))" | 
| 13298 | 93 | by (simp add: OUnion_def) | 
| 12620 | 94 | |
| 95 | lemma OUN_Union_eq: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 96 | "(\<And>x. x \<in> X \<Longrightarrow> Ord(x)) | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 97 | \<Longrightarrow> (\<Union>z < \<Union>(X). C(z)) = (\<Union>x\<in>X. \<Union>z < x. C(z))" | 
| 13298 | 98 | by (simp add: OUnion_def) | 
| 12620 | 99 | |
| 46820 | 100 | (*So that rule_format will get rid of this quantifier...*) | 
| 12763 | 101 | lemma atomize_oall [symmetric, rulify]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 102 | "(\<And>x. x<A \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x<A. P(x))" | 
| 12763 | 103 | by (simp add: oall_def atomize_all atomize_imp) | 
| 104 | ||
| 60770 | 105 | subsubsection \<open>universal quantifier for ordinals\<close> | 
| 13169 | 106 | |
| 107 | lemma oallI [intro!]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 108 | "\<lbrakk>\<And>x. x<A \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> \<forall>x<A. P(x)" | 
| 13298 | 109 | by (simp add: oall_def) | 
| 13169 | 110 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 111 | lemma ospec: "\<lbrakk>\<forall>x<A. P(x); x<A\<rbrakk> \<Longrightarrow> P(x)" | 
| 13298 | 112 | by (simp add: oall_def) | 
| 13169 | 113 | |
| 114 | lemma oallE: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 115 | "\<lbrakk>\<forall>x<A. P(x); P(x) \<Longrightarrow> Q; \<not>x<A \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13298 | 116 | by (simp add: oall_def, blast) | 
| 13169 | 117 | |
| 118 | lemma rev_oallE [elim]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 119 | "\<lbrakk>\<forall>x<A. P(x); \<not>x<A \<Longrightarrow> Q; P(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13298 | 120 | by (simp add: oall_def, blast) | 
| 13169 | 121 | |
| 122 | ||
| 46820 | 123 | (*Trival rewrite rule.  @{term"(\<forall>x<a.P)<->P"} holds only if a is not 0!*)
 | 
| 124 | lemma oall_simp [simp]: "(\<forall>x<a. True) <-> True" | |
| 13170 | 125 | by blast | 
| 13169 | 126 | |
| 127 | (*Congruence rule for rewriting*) | |
| 128 | lemma oall_cong [cong]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 129 | "\<lbrakk>a=a'; \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 130 | \<Longrightarrow> oall(a, \<lambda>x. P(x)) <-> oall(a', \<lambda>x. P'(x))" | 
| 13169 | 131 | by (simp add: oall_def) | 
| 132 | ||
| 133 | ||
| 60770 | 134 | subsubsection \<open>existential quantifier for ordinals\<close> | 
| 13169 | 135 | |
| 136 | lemma oexI [intro]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 137 | "\<lbrakk>P(x); x<A\<rbrakk> \<Longrightarrow> \<exists>x<A. P(x)" | 
| 13298 | 138 | apply (simp add: oex_def, blast) | 
| 13169 | 139 | done | 
| 140 | ||
| 46820 | 141 | (*Not of the general form for such rules... *) | 
| 13169 | 142 | lemma oexCI: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 143 | "\<lbrakk>\<forall>x<A. \<not>P(x) \<Longrightarrow> P(a); a<A\<rbrakk> \<Longrightarrow> \<exists>x<A. P(x)" | 
| 13298 | 144 | apply (simp add: oex_def, blast) | 
| 13169 | 145 | done | 
| 146 | ||
| 147 | lemma oexE [elim!]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 148 | "\<lbrakk>\<exists>x<A. P(x); \<And>x. \<lbrakk>x<A; P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13298 | 149 | apply (simp add: oex_def, blast) | 
| 13169 | 150 | done | 
| 151 | ||
| 152 | lemma oex_cong [cong]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 153 | "\<lbrakk>a=a'; \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 154 | \<Longrightarrow> oex(a, \<lambda>x. P(x)) <-> oex(a', \<lambda>x. P'(x))" | 
| 13169 | 155 | apply (simp add: oex_def cong add: conj_cong) | 
| 156 | done | |
| 157 | ||
| 158 | ||
| 60770 | 159 | subsubsection \<open>Rules for Ordinal-Indexed Unions\<close> | 
| 13169 | 160 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 161 | lemma OUN_I [intro]: "\<lbrakk>a<i; b \<in> B(a)\<rbrakk> \<Longrightarrow> b: (\<Union>z<i. B(z))" | 
| 13170 | 162 | by (unfold OUnion_def lt_def, blast) | 
| 13169 | 163 | |
| 164 | lemma OUN_E [elim!]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 165 | "\<lbrakk>b \<in> (\<Union>z<i. B(z)); \<And>a.\<lbrakk>b \<in> B(a); a<i\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" | 
| 13170 | 166 | apply (unfold OUnion_def lt_def, blast) | 
| 13169 | 167 | done | 
| 168 | ||
| 46820 | 169 | lemma OUN_iff: "b \<in> (\<Union>x<i. B(x)) <-> (\<exists>x<i. b \<in> B(x))" | 
| 13170 | 170 | by (unfold OUnion_def oex_def lt_def, blast) | 
| 13169 | 171 | |
| 172 | lemma OUN_cong [cong]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 173 | "\<lbrakk>i=j; \<And>x. x<j \<Longrightarrow> C(x)=D(x)\<rbrakk> \<Longrightarrow> (\<Union>x<i. C(x)) = (\<Union>x<j. D(x))" | 
| 13169 | 174 | by (simp add: OUnion_def lt_def OUN_iff) | 
| 175 | ||
| 13298 | 176 | lemma lt_induct: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 177 | "\<lbrakk>i<k; \<And>x.\<lbrakk>x<k; \<forall>y<x. P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> P(i)" | 
| 13169 | 178 | apply (simp add: lt_def oall_def) | 
| 13298 | 179 | apply (erule conjE) | 
| 180 | apply (erule Ord_induct, assumption, blast) | |
| 13169 | 181 | done | 
| 182 | ||
| 13253 | 183 | |
| 60770 | 184 | subsection \<open>Quantification over a class\<close> | 
| 13253 | 185 | |
| 24893 | 186 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 187 | "rall" :: "[i\<Rightarrow>o, i\<Rightarrow>o] \<Rightarrow> o" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 188 | "rall(M, P) \<equiv> \<forall>x. M(x) \<longrightarrow> P(x)" | 
| 13253 | 189 | |
| 24893 | 190 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 191 | "rex" :: "[i\<Rightarrow>o, i\<Rightarrow>o] \<Rightarrow> o" where | 
| 76214 | 192 | "rex(M, P) \<equiv> \<exists>x. M(x) \<and> P(x)" | 
| 13253 | 193 | |
| 194 | syntax | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 195 | "_rall" :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o" (\<open>(3\<forall>_[_]./ _)\<close> 10) | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 196 | "_rex" :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o" (\<open>(3\<exists>_[_]./ _)\<close> 10) | 
| 13253 | 197 | translations | 
| 61396 | 198 | "\<forall>x[M]. P" \<rightleftharpoons> "CONST rall(M, \<lambda>x. P)" | 
| 199 | "\<exists>x[M]. P" \<rightleftharpoons> "CONST rex(M, \<lambda>x. P)" | |
| 13253 | 200 | |
| 13298 | 201 | |
| 60770 | 202 | subsubsection\<open>Relativized universal quantifier\<close> | 
| 13253 | 203 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 204 | lemma rallI [intro!]: "\<lbrakk>\<And>x. M(x) \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> \<forall>x[M]. P(x)" | 
| 13253 | 205 | by (simp add: rall_def) | 
| 206 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 207 | lemma rspec: "\<lbrakk>\<forall>x[M]. P(x); M(x)\<rbrakk> \<Longrightarrow> P(x)" | 
| 13253 | 208 | by (simp add: rall_def) | 
| 209 | ||
| 210 | (*Instantiates x first: better for automatic theorem proving?*) | |
| 13298 | 211 | lemma rev_rallE [elim]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 212 | "\<lbrakk>\<forall>x[M]. P(x); \<not> M(x) \<Longrightarrow> Q; P(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13298 | 213 | by (simp add: rall_def, blast) | 
| 13253 | 214 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 215 | lemma rallE: "\<lbrakk>\<forall>x[M]. P(x); P(x) \<Longrightarrow> Q; \<not> M(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13253 | 216 | by blast | 
| 217 | ||
| 61396 | 218 | (*Trival rewrite rule; (\<forall>x[M].P)<->P holds only if A is nonempty!*) | 
| 219 | lemma rall_triv [simp]: "(\<forall>x[M]. P) \<longleftrightarrow> ((\<exists>x. M(x)) \<longrightarrow> P)" | |
| 13253 | 220 | by (simp add: rall_def) | 
| 221 | ||
| 222 | (*Congruence rule for rewriting*) | |
| 223 | lemma rall_cong [cong]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 224 | "(\<And>x. M(x) \<Longrightarrow> P(x) <-> P'(x)) \<Longrightarrow> (\<forall>x[M]. P(x)) <-> (\<forall>x[M]. P'(x))" | 
| 13253 | 225 | by (simp add: rall_def) | 
| 226 | ||
| 13298 | 227 | |
| 60770 | 228 | subsubsection\<open>Relativized existential quantifier\<close> | 
| 13253 | 229 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 230 | lemma rexI [intro]: "\<lbrakk>P(x); M(x)\<rbrakk> \<Longrightarrow> \<exists>x[M]. P(x)" | 
| 13253 | 231 | by (simp add: rex_def, blast) | 
| 232 | ||
| 233 | (*The best argument order when there is only one M(x)*) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 234 | lemma rev_rexI: "\<lbrakk>M(x); P(x)\<rbrakk> \<Longrightarrow> \<exists>x[M]. P(x)" | 
| 13253 | 235 | by blast | 
| 236 | ||
| 46820 | 237 | (*Not of the general form for such rules... *) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 238 | lemma rexCI: "\<lbrakk>\<forall>x[M]. \<not>P(x) \<Longrightarrow> P(a); M(a)\<rbrakk> \<Longrightarrow> \<exists>x[M]. P(x)" | 
| 13253 | 239 | by blast | 
| 240 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 241 | lemma rexE [elim!]: "\<lbrakk>\<exists>x[M]. P(x); \<And>x. \<lbrakk>M(x); P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 13253 | 242 | by (simp add: rex_def, blast) | 
| 243 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 244 | (*We do not even have (\<exists>x[M]. True) <-> True unless A is nonempty\<And>*) | 
| 61396 | 245 | lemma rex_triv [simp]: "(\<exists>x[M]. P) \<longleftrightarrow> ((\<exists>x. M(x)) \<and> P)" | 
| 13253 | 246 | by (simp add: rex_def) | 
| 247 | ||
| 248 | lemma rex_cong [cong]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 249 | "(\<And>x. M(x) \<Longrightarrow> P(x) <-> P'(x)) \<Longrightarrow> (\<exists>x[M]. P(x)) <-> (\<exists>x[M]. P'(x))" | 
| 13253 | 250 | by (simp add: rex_def cong: conj_cong) | 
| 251 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 252 | lemma rall_is_ball [simp]: "(\<forall>x[\<lambda>z. z\<in>A]. P(x)) <-> (\<forall>x\<in>A. P(x))" | 
| 13289 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 253 | by blast | 
| 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 254 | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 255 | lemma rex_is_bex [simp]: "(\<exists>x[\<lambda>z. z\<in>A]. P(x)) <-> (\<exists>x\<in>A. P(x))" | 
| 13289 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 256 | by blast | 
| 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 257 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 258 | lemma atomize_rall: "(\<And>x. M(x) \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x[M]. P(x))" | 
| 13253 | 259 | by (simp add: rall_def atomize_all atomize_imp) | 
| 260 | ||
| 261 | declare atomize_rall [symmetric, rulify] | |
| 262 | ||
| 263 | lemma rall_simps1: | |
| 76214 | 264 | "(\<forall>x[M]. P(x) \<and> Q) <-> (\<forall>x[M]. P(x)) \<and> ((\<forall>x[M]. False) | Q)" | 
| 46820 | 265 | "(\<forall>x[M]. P(x) | Q) <-> ((\<forall>x[M]. P(x)) | Q)" | 
| 266 | "(\<forall>x[M]. P(x) \<longrightarrow> Q) <-> ((\<exists>x[M]. P(x)) \<longrightarrow> Q)" | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 267 | "(\<not>(\<forall>x[M]. P(x))) <-> (\<exists>x[M]. \<not>P(x))" | 
| 13253 | 268 | by blast+ | 
| 269 | ||
| 270 | lemma rall_simps2: | |
| 76214 | 271 | "(\<forall>x[M]. P \<and> Q(x)) <-> ((\<forall>x[M]. False) | P) \<and> (\<forall>x[M]. Q(x))" | 
| 46820 | 272 | "(\<forall>x[M]. P | Q(x)) <-> (P | (\<forall>x[M]. Q(x)))" | 
| 273 | "(\<forall>x[M]. P \<longrightarrow> Q(x)) <-> (P \<longrightarrow> (\<forall>x[M]. Q(x)))" | |
| 13253 | 274 | by blast+ | 
| 275 | ||
| 13289 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 276 | lemmas rall_simps [simp] = rall_simps1 rall_simps2 | 
| 13253 | 277 | |
| 278 | lemma rall_conj_distrib: | |
| 76214 | 279 | "(\<forall>x[M]. P(x) \<and> Q(x)) <-> ((\<forall>x[M]. P(x)) \<and> (\<forall>x[M]. Q(x)))" | 
| 13253 | 280 | by blast | 
| 281 | ||
| 282 | lemma rex_simps1: | |
| 76214 | 283 | "(\<exists>x[M]. P(x) \<and> Q) <-> ((\<exists>x[M]. P(x)) \<and> Q)" | 
| 284 | "(\<exists>x[M]. P(x) | Q) <-> (\<exists>x[M]. P(x)) | ((\<exists>x[M]. True) \<and> Q)" | |
| 285 | "(\<exists>x[M]. P(x) \<longrightarrow> Q) <-> ((\<forall>x[M]. P(x)) \<longrightarrow> ((\<exists>x[M]. True) \<and> Q))" | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
71886diff
changeset | 286 | "(\<not>(\<exists>x[M]. P(x))) <-> (\<forall>x[M]. \<not>P(x))" | 
| 13253 | 287 | by blast+ | 
| 288 | ||
| 289 | lemma rex_simps2: | |
| 76214 | 290 | "(\<exists>x[M]. P \<and> Q(x)) <-> (P \<and> (\<exists>x[M]. Q(x)))" | 
| 291 | "(\<exists>x[M]. P | Q(x)) <-> ((\<exists>x[M]. True) \<and> P) | (\<exists>x[M]. Q(x))" | |
| 46820 | 292 | "(\<exists>x[M]. P \<longrightarrow> Q(x)) <-> (((\<forall>x[M]. False) | P) \<longrightarrow> (\<exists>x[M]. Q(x)))" | 
| 13253 | 293 | by blast+ | 
| 294 | ||
| 13289 
53e201efdaa2
miniscoping for class-bounded quantifiers (rall and rex)
 paulson parents: 
13253diff
changeset | 295 | lemmas rex_simps [simp] = rex_simps1 rex_simps2 | 
| 13253 | 296 | |
| 297 | lemma rex_disj_distrib: | |
| 46820 | 298 | "(\<exists>x[M]. P(x) | Q(x)) <-> ((\<exists>x[M]. P(x)) | (\<exists>x[M]. Q(x)))" | 
| 13253 | 299 | by blast | 
| 300 | ||
| 301 | ||
| 60770 | 302 | subsubsection\<open>One-point rule for bounded quantifiers\<close> | 
| 13253 | 303 | |
| 46820 | 304 | lemma rex_triv_one_point1 [simp]: "(\<exists>x[M]. x=a) <-> ( M(a))" | 
| 13253 | 305 | by blast | 
| 306 | ||
| 46820 | 307 | lemma rex_triv_one_point2 [simp]: "(\<exists>x[M]. a=x) <-> ( M(a))" | 
| 13253 | 308 | by blast | 
| 309 | ||
| 76214 | 310 | lemma rex_one_point1 [simp]: "(\<exists>x[M]. x=a \<and> P(x)) <-> ( M(a) \<and> P(a))" | 
| 13253 | 311 | by blast | 
| 312 | ||
| 76214 | 313 | lemma rex_one_point2 [simp]: "(\<exists>x[M]. a=x \<and> P(x)) <-> ( M(a) \<and> P(a))" | 
| 13253 | 314 | by blast | 
| 315 | ||
| 46820 | 316 | lemma rall_one_point1 [simp]: "(\<forall>x[M]. x=a \<longrightarrow> P(x)) <-> ( M(a) \<longrightarrow> P(a))" | 
| 13253 | 317 | by blast | 
| 318 | ||
| 46820 | 319 | lemma rall_one_point2 [simp]: "(\<forall>x[M]. a=x \<longrightarrow> P(x)) <-> ( M(a) \<longrightarrow> P(a))" | 
| 13253 | 320 | by blast | 
| 321 | ||
| 322 | ||
| 60770 | 323 | subsubsection\<open>Sets as Classes\<close> | 
| 13298 | 324 | |
| 24893 | 325 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 326 | setclass :: "[i,i] \<Rightarrow> o" (\<open>##_\<close> [40] 40) where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 327 | "setclass(A) \<equiv> \<lambda>x. x \<in> A" | 
| 13298 | 328 | |
| 46820 | 329 | lemma setclass_iff [simp]: "setclass(A,x) <-> x \<in> A" | 
| 13362 | 330 | by (simp add: setclass_def) | 
| 13298 | 331 | |
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13615diff
changeset | 332 | lemma rall_setclass_is_ball [simp]: "(\<forall>x[##A]. P(x)) <-> (\<forall>x\<in>A. P(x))" | 
| 13298 | 333 | by auto | 
| 334 | ||
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13615diff
changeset | 335 | lemma rex_setclass_is_bex [simp]: "(\<exists>x[##A]. P(x)) <-> (\<exists>x\<in>A. P(x))" | 
| 13298 | 336 | by auto | 
| 337 | ||
| 338 | ||
| 13169 | 339 | ML | 
| 60770 | 340 | \<open> | 
| 13169 | 341 | val Ord_atomize = | 
| 69593 | 342 |   atomize ([(\<^const_name>\<open>oall\<close>, @{thms ospec}), (\<^const_name>\<open>rall\<close>, @{thms rspec})] @
 | 
| 56250 | 343 | ZF_conn_pairs, ZF_mem_pairs); | 
| 60770 | 344 | \<close> | 
| 345 | declaration \<open>fn _ => | |
| 59647 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59498diff
changeset | 346 | Simplifier.map_ss (Simplifier.set_mksimps (fn ctxt => | 
| 60822 | 347 | map mk_eq o Ord_atomize o Variable.gen_all ctxt)) | 
| 60770 | 348 | \<close> | 
| 13169 | 349 | |
| 60770 | 350 | text \<open>Setting up the one-point-rule simproc\<close> | 
| 13253 | 351 | |
| 76214 | 352 | simproc_setup defined_rex ("\<exists>x[M]. P(x) \<and> Q(x)") = \<open>
 | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
76215diff
changeset | 353 |   K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms rex_def}))
 | 
| 60770 | 354 | \<close> | 
| 13253 | 355 | |
| 60770 | 356 | simproc_setup defined_rall ("\<forall>x[M]. P(x) \<longrightarrow> Q(x)") = \<open>
 | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
76215diff
changeset | 357 |   K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms rall_def}))
 | 
| 60770 | 358 | \<close> | 
| 13253 | 359 | |
| 2469 | 360 | end |