| author | wenzelm | 
| Wed, 11 Sep 2024 23:07:18 +0200 | |
| changeset 80865 | 7c20c207af48 | 
| parent 76216 | 9fc34f76b4e8 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/AC/Hartog.thy | 
| 2 | Author: Krzysztof Grabczewski | |
| 1123 | 3 | |
| 4 | Hartog's function. | |
| 5 | *) | |
| 6 | ||
| 27678 | 7 | theory Hartog | 
| 8 | imports AC_Equiv | |
| 9 | begin | |
| 12776 | 10 | |
| 24893 | 11 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 12 | Hartog :: "i \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 13 | "Hartog(X) \<equiv> \<mu> i. \<not> i \<lesssim> X" | 
| 12776 | 14 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 15 | lemma Ords_in_set: "\<forall>a. Ord(a) \<longrightarrow> a \<in> X \<Longrightarrow> P" | 
| 46471 | 16 | apply (rule_tac X = "{y \<in> X. Ord (y) }" in ON_class [elim_format])
 | 
| 12776 | 17 | apply fast | 
| 18 | done | |
| 1123 | 19 | |
| 12776 | 20 | lemma Ord_lepoll_imp_ex_well_ord: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 21 | "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> | 
| 76214 | 22 | \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> (\<exists>R. well_ord(Y,R) \<and> ordertype(Y,R)=a)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 23 | unfolding lepoll_def | 
| 12776 | 24 | apply (erule exE) | 
| 25 | apply (intro exI conjI) | |
| 26 | apply (erule inj_is_fun [THEN fun_is_rel, THEN image_subset]) | |
| 27 | apply (rule well_ord_rvimage [OF bij_is_inj well_ord_Memrel]) | |
| 28 | apply (erule restrict_bij [THEN bij_converse_bij]) | |
| 12820 | 29 | apply (rule subset_refl, assumption) | 
| 12776 | 30 | apply (rule trans) | 
| 31 | apply (rule bij_ordertype_vimage) | |
| 32 | apply (erule restrict_bij [THEN bij_converse_bij]) | |
| 33 | apply (rule subset_refl) | |
| 34 | apply (erule well_ord_Memrel) | |
| 35 | apply (erule ordertype_Memrel) | |
| 36 | done | |
| 37 | ||
| 38 | lemma Ord_lepoll_imp_eq_ordertype: | |
| 76214 | 39 | "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X \<and> (\<exists>R. R \<subseteq> X*X \<and> ordertype(Y,R)=a)" | 
| 12820 | 40 | apply (drule Ord_lepoll_imp_ex_well_ord, assumption, clarify) | 
| 12776 | 41 | apply (intro exI conjI) | 
| 42 | apply (erule_tac [3] ordertype_Int, auto) | |
| 43 | done | |
| 1123 | 44 | |
| 12776 | 45 | lemma Ords_lepoll_set_lemma: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 46 | "(\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X) \<Longrightarrow> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
46471diff
changeset | 47 | \<forall>a. Ord(a) \<longrightarrow> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 48 |         a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=\<langle>Y,R\<rangle> \<and> ordertype(Y,R)=b}"
 | 
| 12776 | 49 | apply (intro allI impI) | 
| 50 | apply (elim allE impE, assumption) | |
| 51 | apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) | |
| 52 | done | |
| 53 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 54 | lemma Ords_lepoll_set: "\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X \<Longrightarrow> P" | 
| 12776 | 55 | by (erule Ords_lepoll_set_lemma [THEN Ords_in_set]) | 
| 56 | ||
| 76214 | 57 | lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) \<and> \<not>a \<lesssim> X" | 
| 12776 | 58 | apply (rule ccontr) | 
| 59 | apply (best intro: Ords_lepoll_set) | |
| 60 | done | |
| 1123 | 61 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 62 | lemma not_Hartog_lepoll_self: "\<not> Hartog(A) \<lesssim> A" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 63 | unfolding Hartog_def | 
| 12776 | 64 | apply (rule ex_Ord_not_lepoll [THEN exE]) | 
| 65 | apply (rule LeastI, auto) | |
| 66 | done | |
| 67 | ||
| 45602 | 68 | lemmas Hartog_lepoll_selfE = not_Hartog_lepoll_self [THEN notE] | 
| 1123 | 69 | |
| 12776 | 70 | lemma Ord_Hartog: "Ord(Hartog(A))" | 
| 71 | by (unfold Hartog_def, rule Ord_Least) | |
| 72 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 73 | lemma less_HartogE1: "\<lbrakk>i < Hartog(A); \<not> i \<lesssim> A\<rbrakk> \<Longrightarrow> P" | 
| 12776 | 74 | by (unfold Hartog_def, fast elim: less_LeastE) | 
| 75 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 76 | lemma less_HartogE: "\<lbrakk>i < Hartog(A); i \<approx> Hartog(A)\<rbrakk> \<Longrightarrow> P" | 
| 12776 | 77 | by (blast intro: less_HartogE1 eqpoll_sym eqpoll_imp_lepoll | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 78 | lepoll_trans [THEN Hartog_lepoll_selfE]) | 
| 12776 | 79 | |
| 80 | lemma Card_Hartog: "Card(Hartog(A))" | |
| 81 | by (fast intro!: CardI Ord_Hartog elim: less_HartogE) | |
| 1123 | 82 | |
| 1203 | 83 | end |