| author | sultana | 
| Wed, 04 Apr 2012 16:29:16 +0100 | |
| changeset 47358 | 26c4e431ef05 | 
| parent 47308 | 9caab698dbe4 | 
| child 47361 | 87c0eaf04bad | 
| child 47435 | e1b761c216ac | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Quotient.thy | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | Author: Cezary Kaliszyk and Christian Urban | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | |
| 35294 | 5 | header {* Definition of Quotient Types *}
 | 
| 6 | ||
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | theory Quotient | 
| 47308 | 8 | imports Plain Hilbert_Choice Equiv_Relations Lifting | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 9 | keywords | 
| 47308 | 10 | "print_quotmapsQ3" "print_quotientsQ3" "print_quotconsts" :: diag and | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46947diff
changeset | 11 | "quotient_type" :: thy_goal and "/" and | 
| 47091 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 kuncar parents: 
46950diff
changeset | 12 | "quotient_definition" :: thy_goal | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | uses | 
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 14 |   ("Tools/Quotient/quotient_info.ML")
 | 
| 45680 | 15 |   ("Tools/Quotient/quotient_type.ML")
 | 
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 16 |   ("Tools/Quotient/quotient_def.ML")
 | 
| 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 17 |   ("Tools/Quotient/quotient_term.ML")
 | 
| 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 18 |   ("Tools/Quotient/quotient_tacs.ML")
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 19 | begin | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | text {*
 | 
| 45961 | 22 | An aside: contravariant functorial structure of sets. | 
| 23 | *} | |
| 24 | ||
| 25 | enriched_type vimage | |
| 26 | by (simp_all add: fun_eq_iff vimage_compose) | |
| 27 | ||
| 28 | text {*
 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | Basic definition for equivalence relations | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 30 | that are represented by predicates. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | *} | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 32 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 33 | text {* Composition of Relations *}
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 34 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 35 | abbreviation | 
| 40818 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 36 |   rel_conj :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" (infixr "OOO" 75)
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 37 | where | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | "r1 OOO r2 \<equiv> r1 OO r2 OO r1" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 | lemma eq_comp_r: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | shows "((op =) OOO R) = R" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 42 | by (auto simp add: fun_eq_iff) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 43 | |
| 35294 | 44 | subsection {* Respects predicate *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 46 | definition | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 47 |   Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 | where | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 49 |   "Respects R = {x. R x x}"
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | lemma in_respects: | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 52 | shows "x \<in> Respects R \<longleftrightarrow> R x x" | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 53 | unfolding Respects_def by simp | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | |
| 44413 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 55 | subsection {* set map (vimage) and set relation *}
 | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 56 | |
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 57 | definition "set_rel R xs ys \<equiv> \<forall>x y. R x y \<longrightarrow> x \<in> xs \<longleftrightarrow> y \<in> ys" | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 58 | |
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 59 | lemma vimage_id: | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 60 | "vimage id = id" | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 61 | unfolding vimage_def fun_eq_iff by auto | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 62 | |
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 63 | lemma set_rel_eq: | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 64 | "set_rel op = = op =" | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 65 | by (subst fun_eq_iff, subst fun_eq_iff) (simp add: set_eq_iff set_rel_def) | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 66 | |
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 67 | lemma set_rel_equivp: | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 68 | assumes e: "equivp R" | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 69 | shows "set_rel R xs ys \<longleftrightarrow> xs = ys \<and> (\<forall>x y. x \<in> xs \<longrightarrow> R x y \<longrightarrow> y \<in> xs)" | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 70 | unfolding set_rel_def | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 71 | using equivp_reflp[OF e] | 
| 44921 | 72 | by auto (metis, metis equivp_symp[OF e]) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | |
| 35294 | 74 | subsection {* Quotient Predicate *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 76 | definition | 
| 47308 | 77 | "Quotient3 R Abs Rep \<longleftrightarrow> | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 78 | (\<forall>a. Abs (Rep a) = a) \<and> (\<forall>a. R (Rep a) (Rep a)) \<and> | 
| 40818 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 79 | (\<forall>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s)" | 
| 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 80 | |
| 47308 | 81 | lemma Quotient3I: | 
| 40818 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 82 | assumes "\<And>a. Abs (Rep a) = a" | 
| 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 83 | and "\<And>a. R (Rep a) (Rep a)" | 
| 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 84 | and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s" | 
| 47308 | 85 | shows "Quotient3 R Abs Rep" | 
| 86 | using assms unfolding Quotient3_def by blast | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 87 | |
| 47308 | 88 | lemma Quotient3_abs_rep: | 
| 89 | assumes a: "Quotient3 R Abs Rep" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | shows "Abs (Rep a) = a" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | using a | 
| 47308 | 92 | unfolding Quotient3_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 93 | by simp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 94 | |
| 47308 | 95 | lemma Quotient3_rep_reflp: | 
| 96 | assumes a: "Quotient3 R Abs Rep" | |
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 97 | shows "R (Rep a) (Rep a)" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 98 | using a | 
| 47308 | 99 | unfolding Quotient3_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 100 | by blast | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 101 | |
| 47308 | 102 | lemma Quotient3_rel: | 
| 103 | assumes a: "Quotient3 R Abs Rep" | |
| 40818 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 104 |   shows "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 105 | using a | 
| 47308 | 106 | unfolding Quotient3_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 107 | by blast | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 108 | |
| 47308 | 109 | lemma Quotient3_refl1: | 
| 110 | assumes a: "Quotient3 R Abs Rep" | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 111 | shows "R r s \<Longrightarrow> R r r" | 
| 47308 | 112 | using a unfolding Quotient3_def | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 113 | by fast | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 114 | |
| 47308 | 115 | lemma Quotient3_refl2: | 
| 116 | assumes a: "Quotient3 R Abs Rep" | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 117 | shows "R r s \<Longrightarrow> R s s" | 
| 47308 | 118 | using a unfolding Quotient3_def | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 119 | by fast | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 120 | |
| 47308 | 121 | lemma Quotient3_rel_rep: | 
| 122 | assumes a: "Quotient3 R Abs Rep" | |
| 40818 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 haftmann parents: 
40814diff
changeset | 123 | shows "R (Rep a) (Rep b) \<longleftrightarrow> a = b" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 124 | using a | 
| 47308 | 125 | unfolding Quotient3_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 126 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | |
| 47308 | 128 | lemma Quotient3_rep_abs: | 
| 129 | assumes a: "Quotient3 R Abs Rep" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | shows "R r r \<Longrightarrow> R (Rep (Abs r)) r" | 
| 47308 | 131 | using a unfolding Quotient3_def | 
| 132 | by blast | |
| 133 | ||
| 134 | lemma Quotient3_rel_abs: | |
| 135 | assumes a: "Quotient3 R Abs Rep" | |
| 136 | shows "R r s \<Longrightarrow> Abs r = Abs s" | |
| 137 | using a unfolding Quotient3_def | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | by blast | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 139 | |
| 47308 | 140 | lemma Quotient3_symp: | 
| 141 | assumes a: "Quotient3 R Abs Rep" | |
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 142 | shows "symp R" | 
| 47308 | 143 | using a unfolding Quotient3_def using sympI by metis | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | |
| 47308 | 145 | lemma Quotient3_transp: | 
| 146 | assumes a: "Quotient3 R Abs Rep" | |
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 147 | shows "transp R" | 
| 47308 | 148 | using a unfolding Quotient3_def using transpI by (metis (full_types)) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 149 | |
| 47308 | 150 | lemma Quotient3_part_equivp: | 
| 151 | assumes a: "Quotient3 R Abs Rep" | |
| 152 | shows "part_equivp R" | |
| 153 | by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp a part_equivpI) | |
| 154 | ||
| 155 | lemma identity_quotient3: | |
| 156 | shows "Quotient3 (op =) id id" | |
| 157 | unfolding Quotient3_def id_def | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | by blast | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | |
| 47308 | 160 | lemma fun_quotient3: | 
| 161 | assumes q1: "Quotient3 R1 abs1 rep1" | |
| 162 | and q2: "Quotient3 R2 abs2 rep2" | |
| 163 | shows "Quotient3 (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 164 | proof - | 
| 47308 | 165 | have "\<And>a.(rep1 ---> abs2) ((abs1 ---> rep2) a) = a" | 
| 166 | using q1 q2 by (simp add: Quotient3_def fun_eq_iff) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 167 | moreover | 
| 47308 | 168 | have "\<And>a.(R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 169 | by (rule fun_relI) | 
| 47308 | 170 | (insert q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2], | 
| 171 | simp (no_asm) add: Quotient3_def, simp) | |
| 172 | ||
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 173 | moreover | 
| 47308 | 174 |   {
 | 
| 175 | fix r s | |
| 176 | have "(R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and> | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 177 | (rep1 ---> abs2) r = (rep1 ---> abs2) s)" | 
| 47308 | 178 | proof - | 
| 179 | ||
| 180 | have "(R1 ===> R2) r s \<Longrightarrow> (R1 ===> R2) r r" unfolding fun_rel_def | |
| 181 | using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] | |
| 182 | by (metis (full_types) part_equivp_def) | |
| 183 | moreover have "(R1 ===> R2) r s \<Longrightarrow> (R1 ===> R2) s s" unfolding fun_rel_def | |
| 184 | using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] | |
| 185 | by (metis (full_types) part_equivp_def) | |
| 186 | moreover have "(R1 ===> R2) r s \<Longrightarrow> (rep1 ---> abs2) r = (rep1 ---> abs2) s" | |
| 187 | apply(auto simp add: fun_rel_def fun_eq_iff) using q1 q2 unfolding Quotient3_def by metis | |
| 188 | moreover have "((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and> | |
| 189 | (rep1 ---> abs2) r = (rep1 ---> abs2) s) \<Longrightarrow> (R1 ===> R2) r s" | |
| 190 | apply(auto simp add: fun_rel_def fun_eq_iff) using q1 q2 unfolding Quotient3_def | |
| 191 | by (metis map_fun_apply) | |
| 192 | ||
| 193 | ultimately show ?thesis by blast | |
| 194 | qed | |
| 195 | } | |
| 196 | ultimately show ?thesis by (intro Quotient3I) (assumption+) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 197 | qed | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 198 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 199 | lemma abs_o_rep: | 
| 47308 | 200 | assumes a: "Quotient3 R Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 201 | shows "Abs o Rep = id" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 202 | unfolding fun_eq_iff | 
| 47308 | 203 | by (simp add: Quotient3_abs_rep[OF a]) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 204 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 205 | lemma equals_rsp: | 
| 47308 | 206 | assumes q: "Quotient3 R Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 207 | and a: "R xa xb" "R ya yb" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 208 | shows "R xa ya = R xb yb" | 
| 47308 | 209 | using a Quotient3_symp[OF q] Quotient3_transp[OF q] | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 210 | by (blast elim: sympE transpE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 211 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 212 | lemma lambda_prs: | 
| 47308 | 213 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 214 | and q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 215 | shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 216 | unfolding fun_eq_iff | 
| 47308 | 217 | using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 218 | by simp | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 219 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 220 | lemma lambda_prs1: | 
| 47308 | 221 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 222 | and q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 223 | shows "(Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x) = (\<lambda>x. f x)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 224 | unfolding fun_eq_iff | 
| 47308 | 225 | using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 226 | by simp | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 227 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 228 | lemma rep_abs_rsp: | 
| 47308 | 229 | assumes q: "Quotient3 R Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 230 | and a: "R x1 x2" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 231 | shows "R x1 (Rep (Abs x2))" | 
| 47308 | 232 | using a Quotient3_rel[OF q] Quotient3_abs_rep[OF q] Quotient3_rep_reflp[OF q] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 233 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 234 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 235 | lemma rep_abs_rsp_left: | 
| 47308 | 236 | assumes q: "Quotient3 R Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 237 | and a: "R x1 x2" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 238 | shows "R (Rep (Abs x1)) x2" | 
| 47308 | 239 | using a Quotient3_rel[OF q] Quotient3_abs_rep[OF q] Quotient3_rep_reflp[OF q] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 240 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 241 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 242 | text{*
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 243 | In the following theorem R1 can be instantiated with anything, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 244 | but we know some of the types of the Rep and Abs functions; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 245 | so by solving Quotient assumptions we can get a unique R1 that | 
| 35236 
fd8231b16203
quote the constant and theorem name with @{text}
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35222diff
changeset | 246 |   will be provable; which is why we need to use @{text apply_rsp} and
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 247 | not the primed version *} | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 248 | |
| 47308 | 249 | lemma apply_rspQ3: | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 250 | fixes f g::"'a \<Rightarrow> 'c" | 
| 47308 | 251 | assumes q: "Quotient3 R1 Abs1 Rep1" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 252 | and a: "(R1 ===> R2) f g" "R1 x y" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 253 | shows "R2 (f x) (g y)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 254 | using a by (auto elim: fun_relE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 255 | |
| 47308 | 256 | lemma apply_rspQ3'': | 
| 257 | assumes "Quotient3 R Abs Rep" | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 258 | and "(R ===> S) f f" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 259 | shows "S (f (Rep x)) (f (Rep x))" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 260 | proof - | 
| 47308 | 261 | from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient3_rep_reflp) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 262 | then show ?thesis using assms(2) by (auto intro: apply_rsp') | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 263 | qed | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 264 | |
| 35294 | 265 | subsection {* lemmas for regularisation of ball and bex *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 266 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 267 | lemma ball_reg_eqv: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 268 | fixes P :: "'a \<Rightarrow> bool" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 269 | assumes a: "equivp R" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 270 | shows "Ball (Respects R) P = (All P)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 271 | using a | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 272 | unfolding equivp_def | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | by (auto simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 274 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 275 | lemma bex_reg_eqv: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 276 | fixes P :: "'a \<Rightarrow> bool" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 277 | assumes a: "equivp R" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 278 | shows "Bex (Respects R) P = (Ex P)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 279 | using a | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 280 | unfolding equivp_def | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 281 | by (auto simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 282 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 283 | lemma ball_reg_right: | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 284 | assumes a: "\<And>x. x \<in> R \<Longrightarrow> P x \<longrightarrow> Q x" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 285 | shows "All P \<longrightarrow> Ball R Q" | 
| 44921 | 286 | using a by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 287 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 288 | lemma bex_reg_left: | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 289 | assumes a: "\<And>x. x \<in> R \<Longrightarrow> Q x \<longrightarrow> P x" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 290 | shows "Bex R Q \<longrightarrow> Ex P" | 
| 44921 | 291 | using a by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 292 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 293 | lemma ball_reg_left: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 294 | assumes a: "equivp R" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 295 | shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 296 | using a by (metis equivp_reflp in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 297 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 298 | lemma bex_reg_right: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 299 | assumes a: "equivp R" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 300 | shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 301 | using a by (metis equivp_reflp in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 302 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 303 | lemma ball_reg_eqv_range: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 304 | fixes P::"'a \<Rightarrow> bool" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 305 | and x::"'a" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 306 | assumes a: "equivp R2" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 307 | shows "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 308 | apply(rule iffI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 309 | apply(rule allI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 310 | apply(drule_tac x="\<lambda>y. f x" in bspec) | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 311 | apply(simp add: in_respects fun_rel_def) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 312 | apply(rule impI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 313 | using a equivp_reflp_symp_transp[of "R2"] | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 314 | apply (auto elim: equivpE reflpE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 315 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 316 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 317 | lemma bex_reg_eqv_range: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 318 | assumes a: "equivp R2" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 319 | shows "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 320 | apply(auto) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 321 | apply(rule_tac x="\<lambda>y. f x" in bexI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 322 | apply(simp) | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 323 | apply(simp add: Respects_def in_respects fun_rel_def) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 324 | apply(rule impI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 325 | using a equivp_reflp_symp_transp[of "R2"] | 
| 40814 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 haftmann parents: 
40615diff
changeset | 326 | apply (auto elim: equivpE reflpE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 327 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 328 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 329 | (* Next four lemmas are unused *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 330 | lemma all_reg: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 331 | assumes a: "!x :: 'a. (P x --> Q x)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 332 | and b: "All P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 333 | shows "All Q" | 
| 44921 | 334 | using a b by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 335 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 336 | lemma ex_reg: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 337 | assumes a: "!x :: 'a. (P x --> Q x)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 338 | and b: "Ex P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 339 | shows "Ex Q" | 
| 44921 | 340 | using a b by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 341 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 342 | lemma ball_reg: | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 343 | assumes a: "!x :: 'a. (x \<in> R --> P x --> Q x)" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 344 | and b: "Ball R P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 345 | shows "Ball R Q" | 
| 44921 | 346 | using a b by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 347 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 348 | lemma bex_reg: | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 349 | assumes a: "!x :: 'a. (x \<in> R --> P x --> Q x)" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 350 | and b: "Bex R P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 351 | shows "Bex R Q" | 
| 44921 | 352 | using a b by fast | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 353 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 354 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 355 | lemma ball_all_comm: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 356 | assumes "\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 357 | shows "(\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 358 | using assms by auto | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 359 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 360 | lemma bex_ex_comm: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 361 | assumes "(\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 362 | shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 363 | using assms by auto | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 364 | |
| 35294 | 365 | subsection {* Bounded abstraction *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 366 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 367 | definition | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 368 |   Babs :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 369 | where | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 370 | "x \<in> p \<Longrightarrow> Babs p m x = m x" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 371 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 372 | lemma babs_rsp: | 
| 47308 | 373 | assumes q: "Quotient3 R1 Abs1 Rep1" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 374 | and a: "(R1 ===> R2) f g" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 375 | shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 376 | apply (auto simp add: Babs_def in_respects fun_rel_def) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 377 | apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1") | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 378 | using a apply (simp add: Babs_def fun_rel_def) | 
| 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 379 | apply (simp add: in_respects fun_rel_def) | 
| 47308 | 380 | using Quotient3_rel[OF q] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 381 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 382 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 383 | lemma babs_prs: | 
| 47308 | 384 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 385 | and q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 386 | shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 387 | apply (rule ext) | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 388 | apply (simp add:) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 389 | apply (subgoal_tac "Rep1 x \<in> Respects R1") | 
| 47308 | 390 | apply (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]) | 
| 391 | apply (simp add: in_respects Quotient3_rel_rep[OF q1]) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 392 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 393 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 394 | lemma babs_simp: | 
| 47308 | 395 | assumes q: "Quotient3 R1 Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 396 | shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 397 | apply(rule iffI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 398 | apply(simp_all only: babs_rsp[OF q]) | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 399 | apply(auto simp add: Babs_def fun_rel_def) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 400 | apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1") | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 401 | apply(metis Babs_def) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 402 | apply (simp add: in_respects) | 
| 47308 | 403 | using Quotient3_rel[OF q] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 404 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 405 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 406 | (* If a user proves that a particular functional relation | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 407 | is an equivalence this may be useful in regularising *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 408 | lemma babs_reg_eqv: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 409 | shows "equivp R \<Longrightarrow> Babs (Respects R) P = P" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 410 | by (simp add: fun_eq_iff Babs_def in_respects equivp_reflp) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 411 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 412 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 413 | (* 3 lemmas needed for proving repabs_inj *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 414 | lemma ball_rsp: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 415 | assumes a: "(R ===> (op =)) f g" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 416 | shows "Ball (Respects R) f = Ball (Respects R) g" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 417 | using a by (auto simp add: Ball_def in_respects elim: fun_relE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 418 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 419 | lemma bex_rsp: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 420 | assumes a: "(R ===> (op =)) f g" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 421 | shows "(Bex (Respects R) f = Bex (Respects R) g)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 422 | using a by (auto simp add: Bex_def in_respects elim: fun_relE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 423 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 424 | lemma bex1_rsp: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 425 | assumes a: "(R ===> (op =)) f g" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 426 | shows "Ex1 (\<lambda>x. x \<in> Respects R \<and> f x) = Ex1 (\<lambda>x. x \<in> Respects R \<and> g x)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 427 | using a by (auto elim: fun_relE simp add: Ex1_def in_respects) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 428 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 429 | (* 2 lemmas needed for cleaning of quantifiers *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 430 | lemma all_prs: | 
| 47308 | 431 | assumes a: "Quotient3 R absf repf" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 432 | shows "Ball (Respects R) ((absf ---> id) f) = All f" | 
| 47308 | 433 | using a unfolding Quotient3_def Ball_def in_respects id_apply comp_def map_fun_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 434 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 435 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 436 | lemma ex_prs: | 
| 47308 | 437 | assumes a: "Quotient3 R absf repf" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 438 | shows "Bex (Respects R) ((absf ---> id) f) = Ex f" | 
| 47308 | 439 | using a unfolding Quotient3_def Bex_def in_respects id_apply comp_def map_fun_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 440 | by metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 441 | |
| 35294 | 442 | subsection {* @{text Bex1_rel} quantifier *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 443 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 444 | definition | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 445 |   Bex1_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 446 | where | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 447 | "Bex1_rel R P \<longleftrightarrow> (\<exists>x \<in> Respects R. P x) \<and> (\<forall>x \<in> Respects R. \<forall>y \<in> Respects R. ((P x \<and> P y) \<longrightarrow> (R x y)))" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 448 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 449 | lemma bex1_rel_aux: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 450 | "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R x\<rbrakk> \<Longrightarrow> Bex1_rel R y" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 451 | unfolding Bex1_rel_def | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 452 | apply (erule conjE)+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 453 | apply (erule bexE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 454 | apply rule | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 455 | apply (rule_tac x="xa" in bexI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 456 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 457 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 458 | apply rule+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 459 | apply (erule_tac x="xaa" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 460 | prefer 2 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 461 | apply (metis) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 462 | apply (erule_tac x="ya" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 463 | prefer 2 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 464 | apply (metis) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 465 | apply (metis in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 466 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 467 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 468 | lemma bex1_rel_aux2: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 469 | "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R y\<rbrakk> \<Longrightarrow> Bex1_rel R x" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 470 | unfolding Bex1_rel_def | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 471 | apply (erule conjE)+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 472 | apply (erule bexE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 473 | apply rule | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 474 | apply (rule_tac x="xa" in bexI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 475 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 476 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 477 | apply rule+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 478 | apply (erule_tac x="xaa" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 479 | prefer 2 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 480 | apply (metis) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 481 | apply (erule_tac x="ya" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 482 | prefer 2 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 483 | apply (metis) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 484 | apply (metis in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 485 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 486 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 487 | lemma bex1_rel_rsp: | 
| 47308 | 488 | assumes a: "Quotient3 R absf repf" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 489 | shows "((R ===> op =) ===> op =) (Bex1_rel R) (Bex1_rel R)" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 490 | apply (simp add: fun_rel_def) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 491 | apply clarify | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 492 | apply rule | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 493 | apply (simp_all add: bex1_rel_aux bex1_rel_aux2) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 494 | apply (erule bex1_rel_aux2) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 495 | apply assumption | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 496 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 497 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 498 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 499 | lemma ex1_prs: | 
| 47308 | 500 | assumes a: "Quotient3 R absf repf" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 501 | shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f" | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 502 | apply (simp add:) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 503 | apply (subst Bex1_rel_def) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 504 | apply (subst Bex_def) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 505 | apply (subst Ex1_def) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 506 | apply simp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 507 | apply rule | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 508 | apply (erule conjE)+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 509 | apply (erule_tac exE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 510 | apply (erule conjE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 511 | apply (subgoal_tac "\<forall>y. R y y \<longrightarrow> f (absf y) \<longrightarrow> R x y") | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 512 | apply (rule_tac x="absf x" in exI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 513 | apply (simp) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 514 | apply rule+ | 
| 47308 | 515 | using a unfolding Quotient3_def | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 516 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 517 | apply rule+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 518 | apply (erule_tac x="x" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 519 | apply (erule_tac x="y" in ballE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 520 | apply simp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 521 | apply (simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 522 | apply (simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 523 | apply (erule_tac exE) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 524 | apply rule | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 525 | apply (rule_tac x="repf x" in exI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 526 | apply (simp only: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 527 | apply rule | 
| 47308 | 528 | apply (metis Quotient3_rel_rep[OF a]) | 
| 529 | using a unfolding Quotient3_def apply (simp) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 530 | apply rule+ | 
| 47308 | 531 | using a unfolding Quotient3_def in_respects | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 532 | apply metis | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 533 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 534 | |
| 38702 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 535 | lemma bex1_bexeq_reg: | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 536 | shows "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 537 | apply (simp add: Ex1_def Bex1_rel_def in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 538 | apply clarify | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 539 | apply auto | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 540 | apply (rule bexI) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 541 | apply assumption | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 542 | apply (simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 543 | apply (simp add: in_respects) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 544 | apply auto | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 545 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 546 | |
| 38702 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 547 | lemma bex1_bexeq_reg_eqv: | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 548 | assumes a: "equivp R" | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 549 | shows "(\<exists>!x. P x) \<longrightarrow> Bex1_rel R P" | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 550 | using equivp_reflp[OF a] | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 551 | apply (intro impI) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 552 | apply (elim ex1E) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 553 | apply (rule mp[OF bex1_bexeq_reg]) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 554 | apply (rule_tac a="x" in ex1I) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 555 | apply (subst in_respects) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 556 | apply (rule conjI) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 557 | apply assumption | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 558 | apply assumption | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 559 | apply clarify | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 560 | apply (erule_tac x="xa" in allE) | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 561 | apply simp | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 562 | done | 
| 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38317diff
changeset | 563 | |
| 35294 | 564 | subsection {* Various respects and preserve lemmas *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 565 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 566 | lemma quot_rel_rsp: | 
| 47308 | 567 | assumes a: "Quotient3 R Abs Rep" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 568 | shows "(R ===> R ===> op =) R R" | 
| 38317 
cb8e2ac6397b
deleted duplicate lemma
 Christian Urban <urbanc@in.tum.de> parents: 
37986diff
changeset | 569 | apply(rule fun_relI)+ | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 570 | apply(rule equals_rsp[OF a]) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 571 | apply(assumption)+ | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 572 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 573 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 574 | lemma o_prs: | 
| 47308 | 575 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 576 | and q2: "Quotient3 R2 Abs2 Rep2" | |
| 577 | and q3: "Quotient3 R3 Abs3 Rep3" | |
| 36215 
88ff48884d26
respectfullness and preservation of function composition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36123diff
changeset | 578 | shows "((Abs2 ---> Rep3) ---> (Abs1 ---> Rep2) ---> (Rep1 ---> Abs3)) op \<circ> = op \<circ>" | 
| 
88ff48884d26
respectfullness and preservation of function composition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36123diff
changeset | 579 | and "(id ---> (Abs1 ---> id) ---> Rep1 ---> id) op \<circ> = op \<circ>" | 
| 47308 | 580 | using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] Quotient3_abs_rep[OF q3] | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 581 | by (simp_all add: fun_eq_iff) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 582 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 583 | lemma o_rsp: | 
| 36215 
88ff48884d26
respectfullness and preservation of function composition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36123diff
changeset | 584 | "((R2 ===> R3) ===> (R1 ===> R2) ===> (R1 ===> R3)) op \<circ> op \<circ>" | 
| 
88ff48884d26
respectfullness and preservation of function composition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36123diff
changeset | 585 | "(op = ===> (R1 ===> op =) ===> R1 ===> op =) op \<circ> op \<circ>" | 
| 44921 | 586 | by (force elim: fun_relE)+ | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 587 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 588 | lemma cond_prs: | 
| 47308 | 589 | assumes a: "Quotient3 R absf repf" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 590 | shows "absf (if a then repf b else repf c) = (if a then b else c)" | 
| 47308 | 591 | using a unfolding Quotient3_def by auto | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 592 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 593 | lemma if_prs: | 
| 47308 | 594 | assumes q: "Quotient3 R Abs Rep" | 
| 36123 
7f877bbad5b2
add If respectfullness and preservation to Quotient package database
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36116diff
changeset | 595 | shows "(id ---> Rep ---> Rep ---> Abs) If = If" | 
| 47308 | 596 | using Quotient3_abs_rep[OF q] | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 597 | by (auto simp add: fun_eq_iff) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 598 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 599 | lemma if_rsp: | 
| 47308 | 600 | assumes q: "Quotient3 R Abs Rep" | 
| 36123 
7f877bbad5b2
add If respectfullness and preservation to Quotient package database
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36116diff
changeset | 601 | shows "(op = ===> R ===> R ===> R) If If" | 
| 44921 | 602 | by force | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 603 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 604 | lemma let_prs: | 
| 47308 | 605 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 606 | and q2: "Quotient3 R2 Abs2 Rep2" | |
| 37049 
ca1c293e521e
Let rsp and prs in fun_rel/fun_map format
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36276diff
changeset | 607 | shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let" | 
| 47308 | 608 | using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 609 | by (auto simp add: fun_eq_iff) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 610 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 611 | lemma let_rsp: | 
| 37049 
ca1c293e521e
Let rsp and prs in fun_rel/fun_map format
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36276diff
changeset | 612 | shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let" | 
| 44921 | 613 | by (force elim: fun_relE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 614 | |
| 39669 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 615 | lemma id_rsp: | 
| 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 616 | shows "(R ===> R) id id" | 
| 44921 | 617 | by auto | 
| 39669 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 618 | |
| 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 619 | lemma id_prs: | 
| 47308 | 620 | assumes a: "Quotient3 R Abs Rep" | 
| 39669 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 621 | shows "(Rep ---> Abs) id = id" | 
| 47308 | 622 | by (simp add: fun_eq_iff Quotient3_abs_rep [OF a]) | 
| 39669 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 623 | |
| 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
39302diff
changeset | 624 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 625 | locale quot_type = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 626 | fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 627 | and Abs :: "'a set \<Rightarrow> 'b" | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 628 | and Rep :: "'b \<Rightarrow> 'a set" | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37049diff
changeset | 629 | assumes equivp: "part_equivp R" | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 630 | and rep_prop: "\<And>y. \<exists>x. R x x \<and> Rep y = Collect (R x)" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 631 | and rep_inverse: "\<And>x. Abs (Rep x) = x" | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 632 | and abs_inverse: "\<And>c. (\<exists>x. ((R x x) \<and> (c = Collect (R x)))) \<Longrightarrow> (Rep (Abs c)) = c" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 633 | and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 634 | begin | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 635 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 636 | definition | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 637 | abs :: "'a \<Rightarrow> 'b" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 638 | where | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 639 | "abs x = Abs (Collect (R x))" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 640 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 641 | definition | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 642 | rep :: "'b \<Rightarrow> 'a" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 643 | where | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 644 | "rep a = (SOME x. x \<in> Rep a)" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 645 | |
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 646 | lemma some_collect: | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37049diff
changeset | 647 | assumes "R r r" | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 648 | shows "R (SOME x. x \<in> Collect (R r)) = R r" | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 649 | apply simp | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 650 | by (metis assms exE_some equivp[simplified part_equivp_def]) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 651 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 652 | lemma Quotient: | 
| 47308 | 653 | shows "Quotient3 R abs rep" | 
| 654 | unfolding Quotient3_def abs_def rep_def | |
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37049diff
changeset | 655 | proof (intro conjI allI) | 
| 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37049diff
changeset | 656 | fix a r s | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 657 | show x: "R (SOME x. x \<in> Rep a) (SOME x. x \<in> Rep a)" proof - | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 658 | obtain x where r: "R x x" and rep: "Rep a = Collect (R x)" using rep_prop[of a] by auto | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 659 | have "R (SOME x. x \<in> Rep a) x" using r rep some_collect by metis | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 660 | then have "R x (SOME x. x \<in> Rep a)" using part_equivp_symp[OF equivp] by fast | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 661 | then show "R (SOME x. x \<in> Rep a) (SOME x. x \<in> Rep a)" | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 662 | using part_equivp_transp[OF equivp] by (metis `R (SOME x. x \<in> Rep a) x`) | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
37049diff
changeset | 663 | qed | 
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 664 | have "Collect (R (SOME x. x \<in> Rep a)) = (Rep a)" by (metis some_collect rep_prop) | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 665 | then show "Abs (Collect (R (SOME x. x \<in> Rep a))) = a" using rep_inverse by auto | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 666 | have "R r r \<Longrightarrow> R s s \<Longrightarrow> Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> R r = R s" | 
| 44242 | 667 | proof - | 
| 668 | assume "R r r" and "R s s" | |
| 669 | then have "Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> Collect (R r) = Collect (R s)" | |
| 670 | by (metis abs_inverse) | |
| 671 | also have "Collect (R r) = Collect (R s) \<longleftrightarrow> (\<lambda>A x. x \<in> A) (Collect (R r)) = (\<lambda>A x. x \<in> A) (Collect (R s))" | |
| 672 | by rule simp_all | |
| 673 | finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> R r = R s" by simp | |
| 674 | qed | |
| 44204 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 675 | then show "R r s \<longleftrightarrow> R r r \<and> R s s \<and> (Abs (Collect (R r)) = Abs (Collect (R s)))" | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 676 | using equivp[simplified part_equivp_def] by metis | 
| 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
42814diff
changeset | 677 | qed | 
| 44242 | 678 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 679 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 680 | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 681 | subsection {* Quotient composition *}
 | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 682 | |
| 47308 | 683 | lemma OOO_quotient3: | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 684 | fixes R1 :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 685 | fixes Abs1 :: "'a \<Rightarrow> 'b" and Rep1 :: "'b \<Rightarrow> 'a" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 686 | fixes Abs2 :: "'b \<Rightarrow> 'c" and Rep2 :: "'c \<Rightarrow> 'b" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 687 | fixes R2' :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 688 | fixes R2 :: "'b \<Rightarrow> 'b \<Rightarrow> bool" | 
| 47308 | 689 | assumes R1: "Quotient3 R1 Abs1 Rep1" | 
| 690 | assumes R2: "Quotient3 R2 Abs2 Rep2" | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 691 | assumes Abs1: "\<And>x y. R2' x y \<Longrightarrow> R1 x x \<Longrightarrow> R1 y y \<Longrightarrow> R2 (Abs1 x) (Abs1 y)" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 692 | assumes Rep1: "\<And>x y. R2 x y \<Longrightarrow> R2' (Rep1 x) (Rep1 y)" | 
| 47308 | 693 | shows "Quotient3 (R1 OO R2' OO R1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2)" | 
| 694 | apply (rule Quotient3I) | |
| 695 | apply (simp add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 696 | apply simp | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 697 | apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI) | 
| 47308 | 698 | apply (rule Quotient3_rep_reflp [OF R1]) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 699 | apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI [rotated]) | 
| 47308 | 700 | apply (rule Quotient3_rep_reflp [OF R1]) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 701 | apply (rule Rep1) | 
| 47308 | 702 | apply (rule Quotient3_rep_reflp [OF R2]) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 703 | apply safe | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 704 | apply (rename_tac x y) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 705 | apply (drule Abs1) | 
| 47308 | 706 | apply (erule Quotient3_refl2 [OF R1]) | 
| 707 | apply (erule Quotient3_refl1 [OF R1]) | |
| 708 | apply (drule Quotient3_refl1 [OF R2], drule Rep1) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 709 | apply (subgoal_tac "R1 r (Rep1 (Abs1 x))") | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 710 | apply (rule_tac b="Rep1 (Abs1 x)" in pred_compI, assumption) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 711 | apply (erule pred_compI) | 
| 47308 | 712 | apply (erule Quotient3_symp [OF R1, THEN sympD]) | 
| 713 | apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) | |
| 714 | apply (rule conjI, erule Quotient3_refl1 [OF R1]) | |
| 715 | apply (rule conjI, rule Quotient3_rep_reflp [OF R1]) | |
| 716 | apply (subst Quotient3_abs_rep [OF R1]) | |
| 717 | apply (erule Quotient3_rel_abs [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 718 | apply (rename_tac x y) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 719 | apply (drule Abs1) | 
| 47308 | 720 | apply (erule Quotient3_refl2 [OF R1]) | 
| 721 | apply (erule Quotient3_refl1 [OF R1]) | |
| 722 | apply (drule Quotient3_refl2 [OF R2], drule Rep1) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 723 | apply (subgoal_tac "R1 s (Rep1 (Abs1 y))") | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 724 | apply (rule_tac b="Rep1 (Abs1 y)" in pred_compI, assumption) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 725 | apply (erule pred_compI) | 
| 47308 | 726 | apply (erule Quotient3_symp [OF R1, THEN sympD]) | 
| 727 | apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) | |
| 728 | apply (rule conjI, erule Quotient3_refl2 [OF R1]) | |
| 729 | apply (rule conjI, rule Quotient3_rep_reflp [OF R1]) | |
| 730 | apply (subst Quotient3_abs_rep [OF R1]) | |
| 731 | apply (erule Quotient3_rel_abs [OF R1, THEN sym]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 732 | apply simp | 
| 47308 | 733 | apply (rule Quotient3_rel_abs [OF R2]) | 
| 734 | apply (rule Quotient3_rel_abs [OF R1, THEN ssubst], assumption) | |
| 735 | apply (rule Quotient3_rel_abs [OF R1, THEN subst], assumption) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 736 | apply (erule Abs1) | 
| 47308 | 737 | apply (erule Quotient3_refl2 [OF R1]) | 
| 738 | apply (erule Quotient3_refl1 [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 739 | apply (rename_tac a b c d) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 740 | apply simp | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 741 | apply (rule_tac b="Rep1 (Abs1 r)" in pred_compI) | 
| 47308 | 742 | apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) | 
| 743 | apply (rule conjI, erule Quotient3_refl1 [OF R1]) | |
| 744 | apply (simp add: Quotient3_abs_rep [OF R1] Quotient3_rep_reflp [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 745 | apply (rule_tac b="Rep1 (Abs1 s)" in pred_compI [rotated]) | 
| 47308 | 746 | apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) | 
| 747 | apply (simp add: Quotient3_abs_rep [OF R1] Quotient3_rep_reflp [OF R1]) | |
| 748 | apply (erule Quotient3_refl2 [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 749 | apply (rule Rep1) | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 750 | apply (drule Abs1) | 
| 47308 | 751 | apply (erule Quotient3_refl2 [OF R1]) | 
| 752 | apply (erule Quotient3_refl1 [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 753 | apply (drule Abs1) | 
| 47308 | 754 | apply (erule Quotient3_refl2 [OF R1]) | 
| 755 | apply (erule Quotient3_refl1 [OF R1]) | |
| 756 | apply (drule Quotient3_rel_abs [OF R1]) | |
| 757 | apply (drule Quotient3_rel_abs [OF R1]) | |
| 758 | apply (drule Quotient3_rel_abs [OF R1]) | |
| 759 | apply (drule Quotient3_rel_abs [OF R1]) | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 760 | apply simp | 
| 47308 | 761 | apply (rule Quotient3_rel[symmetric, OF R2, THEN iffD2]) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 762 | apply simp | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 763 | done | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 764 | |
| 47308 | 765 | lemma OOO_eq_quotient3: | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 766 | fixes R1 :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 767 | fixes Abs1 :: "'a \<Rightarrow> 'b" and Rep1 :: "'b \<Rightarrow> 'a" | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 768 | fixes Abs2 :: "'b \<Rightarrow> 'c" and Rep2 :: "'c \<Rightarrow> 'b" | 
| 47308 | 769 | assumes R1: "Quotient3 R1 Abs1 Rep1" | 
| 770 | assumes R2: "Quotient3 op= Abs2 Rep2" | |
| 771 | shows "Quotient3 (R1 OOO op=) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2)" | |
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 772 | using assms | 
| 47308 | 773 | by (rule OOO_quotient3) auto | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 774 | |
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 775 | subsection {* Invariant *}
 | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 776 | |
| 47308 | 777 | lemma copy_type_to_Quotient3: | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 778 | assumes "type_definition Rep Abs UNIV" | 
| 47308 | 779 | shows "Quotient3 (op =) Abs Rep" | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 780 | proof - | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 781 | interpret type_definition Rep Abs UNIV by fact | 
| 47308 | 782 | from Abs_inject Rep_inverse show ?thesis by (auto intro!: Quotient3I) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 783 | qed | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 784 | |
| 47308 | 785 | lemma invariant_type_to_Quotient3: | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 786 |   assumes "type_definition Rep Abs {x. P x}"
 | 
| 47308 | 787 | shows "Quotient3 (Lifting.invariant P) Abs Rep" | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 788 | proof - | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 789 |   interpret type_definition Rep Abs "{x. P x}" by fact
 | 
| 47308 | 790 | from Rep Abs_inject Rep_inverse show ?thesis by (auto intro!: Quotient3I simp: invariant_def) | 
| 47096 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 791 | qed | 
| 
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
 kuncar parents: 
47094diff
changeset | 792 | |
| 35294 | 793 | subsection {* ML setup *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 794 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 795 | text {* Auxiliary data for the quotient package *}
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 796 | |
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 797 | use "Tools/Quotient/quotient_info.ML" | 
| 41452 | 798 | setup Quotient_Info.setup | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 799 | |
| 47308 | 800 | declare [[mapQ3 "fun" = (fun_rel, fun_quotient3)]] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 801 | |
| 47308 | 802 | lemmas [quot_thm] = fun_quotient3 | 
| 44553 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 803 | lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp id_rsp | 
| 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 haftmann parents: 
44413diff
changeset | 804 | lemmas [quot_preserve] = if_prs o_prs let_prs id_prs | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 805 | lemmas [quot_equiv] = identity_equivp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 806 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 807 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 808 | text {* Lemmas about simplifying id's. *}
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 809 | lemmas [id_simps] = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 810 | id_def[symmetric] | 
| 40602 | 811 | map_fun_id | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 812 | id_apply | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 813 | id_o | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 814 | o_id | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 815 | eq_comp_r | 
| 44413 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 816 | set_rel_eq | 
| 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
44242diff
changeset | 817 | vimage_id | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 818 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 819 | text {* Translation functions for the lifting process. *}
 | 
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 820 | use "Tools/Quotient/quotient_term.ML" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 821 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 822 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 823 | text {* Definitions of the quotient types. *}
 | 
| 45680 | 824 | use "Tools/Quotient/quotient_type.ML" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 825 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 826 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 827 | text {* Definitions for quotient constants. *}
 | 
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 828 | use "Tools/Quotient/quotient_def.ML" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 829 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 830 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 831 | text {*
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 832 | An auxiliary constant for recording some information | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 833 | about the lifted theorem in a tactic. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 834 | *} | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 835 | definition | 
| 40466 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 836 | Quot_True :: "'a \<Rightarrow> bool" | 
| 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 837 | where | 
| 
c6587375088e
type annotations in specifications; fun_rel_def is no simp rule by default; slightly changed fun_map_def; more on predicates on relation functions; proper HOL equations in definitions
 haftmann parents: 
40031diff
changeset | 838 | "Quot_True x \<longleftrightarrow> True" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 839 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 840 | lemma | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 841 | shows QT_all: "Quot_True (All P) \<Longrightarrow> Quot_True P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 842 | and QT_ex: "Quot_True (Ex P) \<Longrightarrow> Quot_True P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 843 | and QT_ex1: "Quot_True (Ex1 P) \<Longrightarrow> Quot_True P" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 844 | and QT_lam: "Quot_True (\<lambda>x. P x) \<Longrightarrow> (\<And>x. Quot_True (P x))" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 845 | and QT_ext: "(\<And>x. Quot_True (a x) \<Longrightarrow> f x = g x) \<Longrightarrow> (Quot_True a \<Longrightarrow> f = g)" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 846 | by (simp_all add: Quot_True_def ext) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 847 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 848 | lemma QT_imp: "Quot_True a \<equiv> Quot_True b" | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 849 | by (simp add: Quot_True_def) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 850 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 851 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 852 | text {* Tactics for proving the lifted theorems *}
 | 
| 37986 
3b3187adf292
use file names relative to master directory of theory source -- Proof General can now handle that due to the ThyLoad.add_path deception (cf. 3ceccd415145);
 wenzelm parents: 
37593diff
changeset | 853 | use "Tools/Quotient/quotient_tacs.ML" | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 854 | |
| 35294 | 855 | subsection {* Methods / Interface *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 856 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 857 | method_setup lifting = | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 858 |   {* Attrib.thms >> (fn thms => fn ctxt => 
 | 
| 46468 | 859 | SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms)) *} | 
| 42814 | 860 |   {* lift theorems to quotient types *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 861 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 862 | method_setup lifting_setup = | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 863 |   {* Attrib.thm >> (fn thm => fn ctxt => 
 | 
| 46468 | 864 | SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm)) *} | 
| 42814 | 865 |   {* set up the three goals for the quotient lifting procedure *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 866 | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 867 | method_setup descending = | 
| 46468 | 868 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt [])) *}
 | 
| 42814 | 869 |   {* decend theorems to the raw level *}
 | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 870 | |
| 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 871 | method_setup descending_setup = | 
| 46468 | 872 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt [])) *}
 | 
| 42814 | 873 |   {* set up the three goals for the decending theorems *}
 | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37564diff
changeset | 874 | |
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 875 | method_setup partiality_descending = | 
| 46468 | 876 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt [])) *}
 | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 877 |   {* decend theorems to the raw level *}
 | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 878 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 879 | method_setup partiality_descending_setup = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 880 |   {* Scan.succeed (fn ctxt => 
 | 
| 46468 | 881 | SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt [])) *} | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 882 |   {* set up the three goals for the decending theorems *}
 | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45680diff
changeset | 883 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 884 | method_setup regularize = | 
| 46468 | 885 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt)) *}
 | 
| 42814 | 886 |   {* prove the regularization goals from the quotient lifting procedure *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 887 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 888 | method_setup injection = | 
| 46468 | 889 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt)) *}
 | 
| 42814 | 890 |   {* prove the rep/abs injection goals from the quotient lifting procedure *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 891 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 892 | method_setup cleaning = | 
| 46468 | 893 |   {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt)) *}
 | 
| 42814 | 894 |   {* prove the cleaning goals from the quotient lifting procedure *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 895 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 896 | attribute_setup quot_lifted = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 897 |   {* Scan.succeed Quotient_Tacs.lifted_attrib *}
 | 
| 42814 | 898 |   {* lift theorems to quotient types *}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 899 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 900 | no_notation | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 901 | rel_conj (infixr "OOO" 75) and | 
| 40602 | 902 | map_fun (infixr "--->" 55) and | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 903 | fun_rel (infixr "===>" 55) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 904 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 905 | end |