| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 69990 | eb072ce80f82 | 
| child 71262 | a30278c8585f | 
| 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  | 
|
| 60758 | 5  | 
section \<open>Definition of Quotient Types\<close>  | 
| 35294 | 6  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
theory Quotient  | 
| 54555 | 8  | 
imports Lifting  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
9  | 
keywords  | 
| 47308 | 10  | 
"print_quotmapsQ3" "print_quotientsQ3" "print_quotconsts" :: diag and  | 
| 69913 | 11  | 
"quotient_type" :: thy_goal_defn and "/" and  | 
12  | 
"quotient_definition" :: thy_goal_defn  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
13  | 
begin  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
14  | 
|
| 60758 | 15  | 
text \<open>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
16  | 
Basic definition for equivalence relations  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
17  | 
that are represented by predicates.  | 
| 60758 | 18  | 
\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
19  | 
|
| 60758 | 20  | 
text \<open>Composition of Relations\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
22  | 
abbreviation  | 
| 
40818
 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 
haftmann 
parents: 
40814 
diff
changeset
 | 
23  | 
  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
 | 
24  | 
where  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
25  | 
"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
 | 
26  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
27  | 
lemma eq_comp_r:  | 
| 67399 | 28  | 
shows "((=) OOO R) = R"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
29  | 
by (auto simp add: fun_eq_iff)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
30  | 
|
| 63343 | 31  | 
context includes lifting_syntax  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
32  | 
begin  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
33  | 
|
| 60758 | 34  | 
subsection \<open>Quotient Predicate\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
36  | 
definition  | 
| 47308 | 37  | 
"Quotient3 R Abs Rep \<longleftrightarrow>  | 
| 
40814
 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 
haftmann 
parents: 
40615 
diff
changeset
 | 
38  | 
(\<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: 
40814 
diff
changeset
 | 
39  | 
(\<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: 
40814 
diff
changeset
 | 
40  | 
|
| 47308 | 41  | 
lemma Quotient3I:  | 
| 
40818
 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 
haftmann 
parents: 
40814 
diff
changeset
 | 
42  | 
assumes "\<And>a. Abs (Rep a) = a"  | 
| 
 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 
haftmann 
parents: 
40814 
diff
changeset
 | 
43  | 
and "\<And>a. R (Rep a) (Rep a)"  | 
| 
 
b117df72e56b
reorienting iff in Quotient_rel prevents simplifier looping;
 
haftmann 
parents: 
40814 
diff
changeset
 | 
44  | 
and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s"  | 
| 47308 | 45  | 
shows "Quotient3 R Abs Rep"  | 
46  | 
using assms unfolding Quotient3_def by blast  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
47  | 
|
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
48  | 
context  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
49  | 
fixes R Abs Rep  | 
| 47308 | 50  | 
assumes a: "Quotient3 R Abs Rep"  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
51  | 
begin  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
52  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
53  | 
lemma Quotient3_abs_rep:  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
54  | 
"Abs (Rep a) = a"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
55  | 
using a  | 
| 47308 | 56  | 
unfolding Quotient3_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
57  | 
by simp  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
58  | 
|
| 47308 | 59  | 
lemma Quotient3_rep_reflp:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
60  | 
"R (Rep a) (Rep a)"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
61  | 
using a  | 
| 47308 | 62  | 
unfolding Quotient3_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
63  | 
by blast  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
64  | 
|
| 47308 | 65  | 
lemma Quotient3_rel:  | 
| 61799 | 66  | 
"R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" \<comment> \<open>orientation does not loop on rewriting\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
67  | 
using a  | 
| 47308 | 68  | 
unfolding Quotient3_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
69  | 
by blast  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
70  | 
|
| 47308 | 71  | 
lemma Quotient3_refl1:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
72  | 
"R r s \<Longrightarrow> R r r"  | 
| 47308 | 73  | 
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: 
47094 
diff
changeset
 | 
74  | 
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: 
47094 
diff
changeset
 | 
75  | 
|
| 47308 | 76  | 
lemma Quotient3_refl2:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
77  | 
"R r s \<Longrightarrow> R s s"  | 
| 47308 | 78  | 
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: 
47094 
diff
changeset
 | 
79  | 
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: 
47094 
diff
changeset
 | 
80  | 
|
| 47308 | 81  | 
lemma Quotient3_rel_rep:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
82  | 
"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
 | 
83  | 
using a  | 
| 47308 | 84  | 
unfolding Quotient3_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
85  | 
by metis  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
86  | 
|
| 47308 | 87  | 
lemma Quotient3_rep_abs:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
88  | 
"R r r \<Longrightarrow> R (Rep (Abs r)) r"  | 
| 47308 | 89  | 
using a unfolding Quotient3_def  | 
90  | 
by blast  | 
|
91  | 
||
92  | 
lemma Quotient3_rel_abs:  | 
|
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
93  | 
"R r s \<Longrightarrow> Abs r = Abs s"  | 
| 47308 | 94  | 
using a unfolding Quotient3_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
95  | 
by blast  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
96  | 
|
| 47308 | 97  | 
lemma Quotient3_symp:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
98  | 
"symp R"  | 
| 47308 | 99  | 
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
 | 
100  | 
|
| 47308 | 101  | 
lemma Quotient3_transp:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
102  | 
"transp R"  | 
| 47308 | 103  | 
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
 | 
104  | 
|
| 47308 | 105  | 
lemma Quotient3_part_equivp:  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
106  | 
"part_equivp R"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
107  | 
by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp part_equivpI)  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
108  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
109  | 
lemma abs_o_rep:  | 
| 67091 | 110  | 
"Abs \<circ> Rep = id"  | 
| 
54867
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
111  | 
unfolding fun_eq_iff  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
112  | 
by (simp add: Quotient3_abs_rep)  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
113  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
114  | 
lemma equals_rsp:  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
115  | 
assumes b: "R xa xb" "R ya yb"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
116  | 
shows "R xa ya = R xb yb"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
117  | 
using b Quotient3_symp Quotient3_transp  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
118  | 
by (blast elim: sympE transpE)  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
119  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
120  | 
lemma rep_abs_rsp:  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
121  | 
assumes b: "R x1 x2"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
122  | 
shows "R x1 (Rep (Abs x2))"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
123  | 
using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
124  | 
by metis  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
125  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
126  | 
lemma rep_abs_rsp_left:  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
127  | 
assumes b: "R x1 x2"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
128  | 
shows "R (Rep (Abs x1)) x2"  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
129  | 
using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
130  | 
by metis  | 
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
131  | 
|
| 
 
c21a2465cac1
prefer ephemeral interpretation over interpretation in proof contexts;
 
haftmann 
parents: 
54555 
diff
changeset
 | 
132  | 
end  | 
| 47308 | 133  | 
|
134  | 
lemma identity_quotient3:  | 
|
| 67399 | 135  | 
"Quotient3 (=) id id"  | 
| 47308 | 136  | 
unfolding Quotient3_def id_def  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
137  | 
by blast  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
138  | 
|
| 47308 | 139  | 
lemma fun_quotient3:  | 
140  | 
assumes q1: "Quotient3 R1 abs1 rep1"  | 
|
141  | 
and q2: "Quotient3 R2 abs2 rep2"  | 
|
142  | 
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
 | 
143  | 
proof -  | 
| 69990 | 144  | 
have "(rep1 ---> abs2) ((abs1 ---> rep2) a) = a" for a  | 
| 47308 | 145  | 
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
 | 
146  | 
moreover  | 
| 69990 | 147  | 
have "(R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" for a  | 
| 55945 | 148  | 
by (rule rel_funI)  | 
| 47308 | 149  | 
(insert q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2],  | 
150  | 
simp (no_asm) add: Quotient3_def, simp)  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
151  | 
moreover  | 
| 47308 | 152  | 
have "(R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and>  | 
| 69990 | 153  | 
(rep1 ---> abs2) r = (rep1 ---> abs2) s)" for r s  | 
| 47308 | 154  | 
proof -  | 
| 55945 | 155  | 
have "(R1 ===> R2) r s \<Longrightarrow> (R1 ===> R2) r r" unfolding rel_fun_def  | 
| 47308 | 156  | 
using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2]  | 
157  | 
by (metis (full_types) part_equivp_def)  | 
|
| 55945 | 158  | 
moreover have "(R1 ===> R2) r s \<Longrightarrow> (R1 ===> R2) s s" unfolding rel_fun_def  | 
| 47308 | 159  | 
using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2]  | 
160  | 
by (metis (full_types) part_equivp_def)  | 
|
161  | 
moreover have "(R1 ===> R2) r s \<Longrightarrow> (rep1 ---> abs2) r = (rep1 ---> abs2) s"  | 
|
| 69990 | 162  | 
by (auto simp add: rel_fun_def fun_eq_iff)  | 
163  | 
(use q1 q2 in \<open>unfold Quotient3_def, metis\<close>)  | 
|
| 47308 | 164  | 
moreover have "((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and>  | 
165  | 
(rep1 ---> abs2) r = (rep1 ---> abs2) s) \<Longrightarrow> (R1 ===> R2) r s"  | 
|
| 69990 | 166  | 
by (auto simp add: rel_fun_def fun_eq_iff)  | 
167  | 
(use q1 q2 in \<open>unfold Quotient3_def, metis map_fun_apply\<close>)  | 
|
| 47308 | 168  | 
ultimately show ?thesis by blast  | 
| 69990 | 169  | 
qed  | 
170  | 
ultimately show ?thesis by (intro Quotient3I) (assumption+)  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
171  | 
qed  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
173  | 
lemma lambda_prs:  | 
| 47308 | 174  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
175  | 
and q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
176  | 
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: 
39198 
diff
changeset
 | 
177  | 
unfolding fun_eq_iff  | 
| 47308 | 178  | 
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: 
40615 
diff
changeset
 | 
179  | 
by simp  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
180  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
181  | 
lemma lambda_prs1:  | 
| 47308 | 182  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
183  | 
and q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
184  | 
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: 
39198 
diff
changeset
 | 
185  | 
unfolding fun_eq_iff  | 
| 47308 | 186  | 
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: 
40615 
diff
changeset
 | 
187  | 
by simp  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
188  | 
|
| 60758 | 189  | 
text\<open>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
190  | 
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
 | 
191  | 
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
 | 
192  | 
so by solving Quotient assumptions we can get a unique R1 that  | 
| 61799 | 193  | 
will be provable; which is why we need to use \<open>apply_rsp\<close> and  | 
| 60758 | 194  | 
not the primed version\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
195  | 
|
| 47308 | 196  | 
lemma apply_rspQ3:  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
197  | 
fixes f g::"'a \<Rightarrow> 'c"  | 
| 47308 | 198  | 
assumes q: "Quotient3 R1 Abs1 Rep1"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
199  | 
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
 | 
200  | 
shows "R2 (f x) (g y)"  | 
| 55945 | 201  | 
using a by (auto elim: rel_funE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
202  | 
|
| 47308 | 203  | 
lemma apply_rspQ3'':  | 
204  | 
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: 
47094 
diff
changeset
 | 
205  | 
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: 
47094 
diff
changeset
 | 
206  | 
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: 
47094 
diff
changeset
 | 
207  | 
proof -  | 
| 47308 | 208  | 
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: 
47094 
diff
changeset
 | 
209  | 
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: 
47094 
diff
changeset
 | 
210  | 
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: 
47094 
diff
changeset
 | 
211  | 
|
| 60758 | 212  | 
subsection \<open>lemmas for regularisation of ball and bex\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
214  | 
lemma ball_reg_eqv:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
215  | 
fixes P :: "'a \<Rightarrow> bool"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
216  | 
assumes a: "equivp R"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
217  | 
shows "Ball (Respects R) P = (All P)"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
218  | 
using a  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
219  | 
unfolding equivp_def  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
220  | 
by (auto simp add: in_respects)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
222  | 
lemma bex_reg_eqv:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
223  | 
fixes P :: "'a \<Rightarrow> bool"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
224  | 
assumes a: "equivp R"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
225  | 
shows "Bex (Respects R) P = (Ex P)"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
226  | 
using a  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
227  | 
unfolding equivp_def  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
228  | 
by (auto simp add: in_respects)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
230  | 
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: 
44413 
diff
changeset
 | 
231  | 
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
 | 
232  | 
shows "All P \<longrightarrow> Ball R Q"  | 
| 44921 | 233  | 
using a by fast  | 
| 
35222
 
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 bex_reg_left:  | 
| 
44553
 
4d39b032a021
avoid intermixing set and predicates; dropped lemmas mem_rsp and mem_prs (now in Quotient_Set.thy)
 
haftmann 
parents: 
44413 
diff
changeset
 | 
236  | 
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
 | 
237  | 
shows "Bex R Q \<longrightarrow> Ex P"  | 
| 44921 | 238  | 
using a by fast  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
239  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
240  | 
lemma ball_reg_left:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
241  | 
assumes a: "equivp R"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
242  | 
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
 | 
243  | 
using a by (metis equivp_reflp in_respects)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
245  | 
lemma bex_reg_right:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
246  | 
assumes a: "equivp R"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
247  | 
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
 | 
248  | 
using a by (metis equivp_reflp in_respects)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
249  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
250  | 
lemma ball_reg_eqv_range:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
251  | 
fixes P::"'a \<Rightarrow> bool"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
252  | 
and x::"'a"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
253  | 
assumes a: "equivp R2"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
254  | 
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
 | 
255  | 
apply(rule iffI)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
256  | 
apply(rule allI)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
257  | 
apply(drule_tac x="\<lambda>y. f x" in bspec)  | 
| 55945 | 258  | 
apply(simp add: in_respects rel_fun_def)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
259  | 
apply(rule impI)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
260  | 
using a equivp_reflp_symp_transp[of "R2"]  | 
| 
40814
 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 
haftmann 
parents: 
40615 
diff
changeset
 | 
261  | 
apply (auto elim: equivpE reflpE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
262  | 
done  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
263  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
264  | 
lemma bex_reg_eqv_range:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
265  | 
assumes a: "equivp R2"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
266  | 
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
 | 
267  | 
apply(auto)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
268  | 
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
 | 
269  | 
apply(simp)  | 
| 55945 | 270  | 
apply(simp add: Respects_def in_respects rel_fun_def)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
271  | 
apply(rule impI)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
272  | 
using a equivp_reflp_symp_transp[of "R2"]  | 
| 
40814
 
fa64f6278568
moved generic definitions about (partial) equivalence relations from Quotient to Equiv_Relations;
 
haftmann 
parents: 
40615 
diff
changeset
 | 
273  | 
apply (auto elim: equivpE reflpE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
274  | 
done  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
275  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
276  | 
(* Next four lemmas are unused *)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
277  | 
lemma all_reg:  | 
| 67091 | 278  | 
assumes a: "\<forall>x :: 'a. (P x \<longrightarrow> Q x)"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
279  | 
and b: "All P"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
280  | 
shows "All Q"  | 
| 44921 | 281  | 
using a b by fast  | 
| 
35222
 
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 ex_reg:  | 
| 67091 | 284  | 
assumes a: "\<forall>x :: 'a. (P x \<longrightarrow> Q x)"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
285  | 
and b: "Ex P"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
286  | 
shows "Ex Q"  | 
| 44921 | 287  | 
using a b by fast  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
289  | 
lemma ball_reg:  | 
| 67091 | 290  | 
assumes a: "\<forall>x :: 'a. (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
 | 
291  | 
and b: "Ball R P"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
292  | 
shows "Ball R Q"  | 
| 44921 | 293  | 
using a b by fast  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
294  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
295  | 
lemma bex_reg:  | 
| 67091 | 296  | 
assumes a: "\<forall>x :: 'a. (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
 | 
297  | 
and b: "Bex R P"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
298  | 
shows "Bex R Q"  | 
| 44921 | 299  | 
using a b by fast  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
301  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
302  | 
lemma ball_all_comm:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
303  | 
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
 | 
304  | 
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
 | 
305  | 
using assms by auto  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
307  | 
lemma bex_ex_comm:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
308  | 
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
 | 
309  | 
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
 | 
310  | 
using assms by auto  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
311  | 
|
| 60758 | 312  | 
subsection \<open>Bounded abstraction\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
313  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
314  | 
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: 
40031 
diff
changeset
 | 
315  | 
  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
 | 
316  | 
where  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
317  | 
"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
 | 
318  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
319  | 
lemma babs_rsp:  | 
| 47308 | 320  | 
assumes q: "Quotient3 R1 Abs1 Rep1"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
321  | 
and a: "(R1 ===> R2) f g"  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
322  | 
shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)"  | 
| 55945 | 323  | 
apply (auto simp add: Babs_def in_respects rel_fun_def)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
324  | 
apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")  | 
| 55945 | 325  | 
using a apply (simp add: Babs_def rel_fun_def)  | 
326  | 
apply (simp add: in_respects rel_fun_def)  | 
|
| 47308 | 327  | 
using Quotient3_rel[OF q]  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
328  | 
by metis  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
329  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
330  | 
lemma babs_prs:  | 
| 47308 | 331  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
332  | 
and q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
333  | 
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
 | 
334  | 
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: 
40031 
diff
changeset
 | 
335  | 
apply (simp add:)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
336  | 
apply (subgoal_tac "Rep1 x \<in> Respects R1")  | 
| 47308 | 337  | 
apply (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])  | 
338  | 
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
 | 
339  | 
done  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
340  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
341  | 
lemma babs_simp:  | 
| 47308 | 342  | 
assumes q: "Quotient3 R1 Abs Rep"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
343  | 
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
 | 
344  | 
apply(rule iffI)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
345  | 
apply(simp_all only: babs_rsp[OF q])  | 
| 55945 | 346  | 
apply(auto simp add: Babs_def rel_fun_def)  | 
| 68615 | 347  | 
apply(metis Babs_def in_respects Quotient3_rel[OF q])  | 
348  | 
done  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
349  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
350  | 
(* 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
 | 
351  | 
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
 | 
352  | 
lemma babs_reg_eqv:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
353  | 
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: 
39198 
diff
changeset
 | 
354  | 
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
 | 
355  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
356  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
357  | 
(* 3 lemmas needed for proving repabs_inj *)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
358  | 
lemma ball_rsp:  | 
| 67399 | 359  | 
assumes a: "(R ===> (=)) f g"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
360  | 
shows "Ball (Respects R) f = Ball (Respects R) g"  | 
| 55945 | 361  | 
using a by (auto simp add: Ball_def in_respects elim: rel_funE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
362  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
363  | 
lemma bex_rsp:  | 
| 67399 | 364  | 
assumes a: "(R ===> (=)) f g"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
365  | 
shows "(Bex (Respects R) f = Bex (Respects R) g)"  | 
| 55945 | 366  | 
using a by (auto simp add: Bex_def in_respects elim: rel_funE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
367  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
368  | 
lemma bex1_rsp:  | 
| 67399 | 369  | 
assumes a: "(R ===> (=)) f g"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
370  | 
shows "Ex1 (\<lambda>x. x \<in> Respects R \<and> f x) = Ex1 (\<lambda>x. x \<in> Respects R \<and> g x)"  | 
| 55945 | 371  | 
using a by (auto elim: rel_funE simp add: Ex1_def in_respects)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
372  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
373  | 
(* 2 lemmas needed for cleaning of quantifiers *)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
374  | 
lemma all_prs:  | 
| 47308 | 375  | 
assumes a: "Quotient3 R absf repf"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
376  | 
shows "Ball (Respects R) ((absf ---> id) f) = All f"  | 
| 47308 | 377  | 
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
 | 
378  | 
by metis  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
379  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
380  | 
lemma ex_prs:  | 
| 47308 | 381  | 
assumes a: "Quotient3 R absf repf"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
382  | 
shows "Bex (Respects R) ((absf ---> id) f) = Ex f"  | 
| 47308 | 383  | 
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
 | 
384  | 
by metis  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
385  | 
|
| 61799 | 386  | 
subsection \<open>\<open>Bex1_rel\<close> quantifier\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
387  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
388  | 
definition  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
389  | 
  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
 | 
390  | 
where  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
391  | 
"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
 | 
392  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
393  | 
lemma bex1_rel_aux:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
394  | 
"\<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
 | 
395  | 
unfolding Bex1_rel_def  | 
| 68615 | 396  | 
by (metis in_respects)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
397  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
398  | 
lemma bex1_rel_aux2:  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
399  | 
"\<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
 | 
400  | 
unfolding Bex1_rel_def  | 
| 68615 | 401  | 
by (metis in_respects)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
402  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
403  | 
lemma bex1_rel_rsp:  | 
| 47308 | 404  | 
assumes a: "Quotient3 R absf repf"  | 
| 67399 | 405  | 
shows "((R ===> (=)) ===> (=)) (Bex1_rel R) (Bex1_rel R)"  | 
| 68615 | 406  | 
unfolding rel_fun_def by (metis bex1_rel_aux bex1_rel_aux2)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
407  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
408  | 
lemma ex1_prs:  | 
| 68616 | 409  | 
assumes "Quotient3 R absf repf"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
410  | 
shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f"  | 
| 68616 | 411  | 
apply (auto simp: Bex1_rel_def Respects_def)  | 
412  | 
apply (metis Quotient3_def assms)  | 
|
413  | 
apply (metis (full_types) Quotient3_def assms)  | 
|
414  | 
by (meson Quotient3_rel assms)  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
415  | 
|
| 
38702
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
416  | 
lemma bex1_bexeq_reg:  | 
| 
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
417  | 
shows "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"  | 
| 
56073
 
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
 
nipkow 
parents: 
55945 
diff
changeset
 | 
418  | 
by (auto simp add: Ex1_def Bex1_rel_def Bex_def Ball_def in_respects)  | 
| 68616 | 419  | 
|
| 
38702
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
420  | 
lemma bex1_bexeq_reg_eqv:  | 
| 
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
421  | 
assumes a: "equivp R"  | 
| 
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
422  | 
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: 
38317 
diff
changeset
 | 
423  | 
using equivp_reflp[OF a]  | 
| 68616 | 424  | 
by (metis (full_types) Bex1_rel_def in_respects)  | 
| 
38702
 
72fd257f4343
Quotient Package / lemma for regularization of bex1_rel for equivalence relations
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
38317 
diff
changeset
 | 
425  | 
|
| 60758 | 426  | 
subsection \<open>Various respects and preserve lemmas\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
427  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
428  | 
lemma quot_rel_rsp:  | 
| 47308 | 429  | 
assumes a: "Quotient3 R Abs Rep"  | 
| 67399 | 430  | 
shows "(R ===> R ===> (=)) R R"  | 
| 55945 | 431  | 
apply(rule rel_funI)+  | 
| 68616 | 432  | 
by (meson assms equals_rsp)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
433  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
434  | 
lemma o_prs:  | 
| 47308 | 435  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
436  | 
and q2: "Quotient3 R2 Abs2 Rep2"  | 
|
437  | 
and q3: "Quotient3 R3 Abs3 Rep3"  | 
|
| 67399 | 438  | 
shows "((Abs2 ---> Rep3) ---> (Abs1 ---> Rep2) ---> (Rep1 ---> Abs3)) (\<circ>) = (\<circ>)"  | 
439  | 
and "(id ---> (Abs1 ---> id) ---> Rep1 ---> id) (\<circ>) = (\<circ>)"  | 
|
| 47308 | 440  | 
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: 
40031 
diff
changeset
 | 
441  | 
by (simp_all add: fun_eq_iff)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
442  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
443  | 
lemma o_rsp:  | 
| 67399 | 444  | 
"((R2 ===> R3) ===> (R1 ===> R2) ===> (R1 ===> R3)) (\<circ>) (\<circ>)"  | 
445  | 
"((=) ===> (R1 ===> (=)) ===> R1 ===> (=)) (\<circ>) (\<circ>)"  | 
|
| 55945 | 446  | 
by (force elim: rel_funE)+  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
447  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
448  | 
lemma cond_prs:  | 
| 47308 | 449  | 
assumes a: "Quotient3 R absf repf"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
450  | 
shows "absf (if a then repf b else repf c) = (if a then b else c)"  | 
| 47308 | 451  | 
using a unfolding Quotient3_def by auto  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
452  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
453  | 
lemma if_prs:  | 
| 47308 | 454  | 
assumes q: "Quotient3 R Abs Rep"  | 
| 
36123
 
7f877bbad5b2
add If respectfullness and preservation to Quotient package database
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
36116 
diff
changeset
 | 
455  | 
shows "(id ---> Rep ---> Rep ---> Abs) If = If"  | 
| 47308 | 456  | 
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: 
39198 
diff
changeset
 | 
457  | 
by (auto simp add: fun_eq_iff)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
459  | 
lemma if_rsp:  | 
| 47308 | 460  | 
assumes q: "Quotient3 R Abs Rep"  | 
| 67399 | 461  | 
shows "((=) ===> R ===> R ===> R) If If"  | 
| 44921 | 462  | 
by force  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
463  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
464  | 
lemma let_prs:  | 
| 47308 | 465  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
466  | 
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: 
36276 
diff
changeset
 | 
467  | 
shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let"  | 
| 47308 | 468  | 
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: 
39198 
diff
changeset
 | 
469  | 
by (auto simp add: fun_eq_iff)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
471  | 
lemma let_rsp:  | 
| 
37049
 
ca1c293e521e
Let rsp and prs in fun_rel/fun_map format
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
36276 
diff
changeset
 | 
472  | 
shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let"  | 
| 55945 | 473  | 
by (force elim: rel_funE)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
474  | 
|
| 
39669
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
475  | 
lemma id_rsp:  | 
| 
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
476  | 
shows "(R ===> R) id id"  | 
| 44921 | 477  | 
by auto  | 
| 
39669
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
478  | 
|
| 
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
479  | 
lemma id_prs:  | 
| 47308 | 480  | 
assumes a: "Quotient3 R Abs Rep"  | 
| 
39669
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
481  | 
shows "(Rep ---> Abs) id = id"  | 
| 47308 | 482  | 
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: 
39302 
diff
changeset
 | 
483  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
484  | 
end  | 
| 
39669
 
9e3b035841e4
quotient package: respectfulness and preservation of identity.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
39302 
diff
changeset
 | 
485  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
486  | 
locale quot_type =  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
487  | 
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: 
42814 
diff
changeset
 | 
488  | 
and Abs :: "'a set \<Rightarrow> 'b"  | 
| 
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
489  | 
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: 
37049 
diff
changeset
 | 
490  | 
assumes equivp: "part_equivp R"  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
491  | 
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
 | 
492  | 
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: 
42814 
diff
changeset
 | 
493  | 
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
 | 
494  | 
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
 | 
495  | 
begin  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
496  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
497  | 
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: 
40031 
diff
changeset
 | 
498  | 
abs :: "'a \<Rightarrow> 'b"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
499  | 
where  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
500  | 
"abs x = Abs (Collect (R x))"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
501  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
502  | 
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: 
40031 
diff
changeset
 | 
503  | 
rep :: "'b \<Rightarrow> 'a"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
504  | 
where  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
505  | 
"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
 | 
506  | 
|
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
507  | 
lemma some_collect:  | 
| 
37493
 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
37049 
diff
changeset
 | 
508  | 
assumes "R r r"  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
509  | 
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: 
42814 
diff
changeset
 | 
510  | 
apply simp  | 
| 
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
511  | 
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
 | 
512  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
513  | 
lemma Quotient:  | 
| 47308 | 514  | 
shows "Quotient3 R abs rep"  | 
515  | 
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: 
37049 
diff
changeset
 | 
516  | 
proof (intro conjI allI)  | 
| 
 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
37049 
diff
changeset
 | 
517  | 
fix a r s  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
518  | 
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: 
42814 
diff
changeset
 | 
519  | 
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: 
42814 
diff
changeset
 | 
520  | 
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: 
42814 
diff
changeset
 | 
521  | 
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: 
42814 
diff
changeset
 | 
522  | 
then show "R (SOME x. x \<in> Rep a) (SOME x. x \<in> Rep a)"  | 
| 60758 | 523  | 
using part_equivp_transp[OF equivp] by (metis \<open>R (SOME x. x \<in> Rep a) x\<close>)  | 
| 
37493
 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
37049 
diff
changeset
 | 
524  | 
qed  | 
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
525  | 
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: 
42814 
diff
changeset
 | 
526  | 
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: 
42814 
diff
changeset
 | 
527  | 
have "R r r \<Longrightarrow> R s s \<Longrightarrow> Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> R r = R s"  | 
| 44242 | 528  | 
proof -  | 
529  | 
assume "R r r" and "R s s"  | 
|
530  | 
then have "Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> Collect (R r) = Collect (R s)"  | 
|
531  | 
by (metis abs_inverse)  | 
|
532  | 
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))"  | 
|
533  | 
by rule simp_all  | 
|
534  | 
finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \<longleftrightarrow> R r = R s" by simp  | 
|
535  | 
qed  | 
|
| 
44204
 
3cdc4176638c
Quotient Package: make quotient_type work with separate set type
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
42814 
diff
changeset
 | 
536  | 
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: 
42814 
diff
changeset
 | 
537  | 
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: 
42814 
diff
changeset
 | 
538  | 
qed  | 
| 44242 | 539  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
540  | 
end  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
541  | 
|
| 60758 | 542  | 
subsection \<open>Quotient composition\<close>  | 
| 
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: 
47094 
diff
changeset
 | 
543  | 
|
| 68616 | 544  | 
|
| 47308 | 545  | 
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: 
47094 
diff
changeset
 | 
546  | 
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: 
47094 
diff
changeset
 | 
547  | 
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: 
47094 
diff
changeset
 | 
548  | 
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: 
47094 
diff
changeset
 | 
549  | 
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: 
47094 
diff
changeset
 | 
550  | 
fixes R2 :: "'b \<Rightarrow> 'b \<Rightarrow> bool"  | 
| 47308 | 551  | 
assumes R1: "Quotient3 R1 Abs1 Rep1"  | 
552  | 
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: 
47094 
diff
changeset
 | 
553  | 
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: 
47094 
diff
changeset
 | 
554  | 
assumes Rep1: "\<And>x y. R2 x y \<Longrightarrow> R2' (Rep1 x) (Rep1 y)"  | 
| 47308 | 555  | 
shows "Quotient3 (R1 OO R2' OO R1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2)"  | 
| 68616 | 556  | 
proof -  | 
557  | 
have *: "(R1 OOO R2') r r \<and> (R1 OOO R2') s s \<and> (Abs2 \<circ> Abs1) r = (Abs2 \<circ> Abs1) s  | 
|
558  | 
\<longleftrightarrow> (R1 OOO R2') r s" for r s  | 
|
559  | 
apply safe  | 
|
560  | 
subgoal for a b c d  | 
|
561  | 
apply simp  | 
|
562  | 
apply (rule_tac b="Rep1 (Abs1 r)" in relcomppI)  | 
|
563  | 
using Quotient3_refl1 R1 rep_abs_rsp apply fastforce  | 
|
564  | 
apply (rule_tac b="Rep1 (Abs1 s)" in relcomppI)  | 
|
565  | 
apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2 Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1])  | 
|
566  | 
by (metis Quotient3_rel R1 rep_abs_rsp_left)  | 
|
567  | 
subgoal for x y  | 
|
568  | 
apply (drule Abs1)  | 
|
569  | 
apply (erule Quotient3_refl2 [OF R1])  | 
|
570  | 
apply (erule Quotient3_refl1 [OF R1])  | 
|
571  | 
apply (drule Quotient3_refl1 [OF R2], drule Rep1)  | 
|
572  | 
by (metis (full_types) Quotient3_def R1 relcompp.relcompI)  | 
|
573  | 
subgoal for x y  | 
|
574  | 
apply (drule Abs1)  | 
|
575  | 
apply (erule Quotient3_refl2 [OF R1])  | 
|
576  | 
apply (erule Quotient3_refl1 [OF R1])  | 
|
577  | 
apply (drule Quotient3_refl2 [OF R2], drule Rep1)  | 
|
578  | 
by (metis (full_types) Quotient3_def R1 relcompp.relcompI)  | 
|
579  | 
subgoal for x y  | 
|
580  | 
by simp (metis (full_types) Abs1 Quotient3_rel R1 R2)  | 
|
581  | 
done  | 
|
582  | 
show ?thesis  | 
|
583  | 
apply (rule Quotient3I)  | 
|
584  | 
using * apply (simp_all add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1])  | 
|
585  | 
apply (metis Quotient3_rep_reflp R1 R2 Rep1 relcompp.relcompI)  | 
|
586  | 
done  | 
|
587  | 
qed  | 
|
| 
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: 
47094 
diff
changeset
 | 
588  | 
|
| 47308 | 589  | 
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: 
47094 
diff
changeset
 | 
590  | 
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: 
47094 
diff
changeset
 | 
591  | 
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: 
47094 
diff
changeset
 | 
592  | 
fixes Abs2 :: "'b \<Rightarrow> 'c" and Rep2 :: "'c \<Rightarrow> 'b"  | 
| 47308 | 593  | 
assumes R1: "Quotient3 R1 Abs1 Rep1"  | 
| 67399 | 594  | 
assumes R2: "Quotient3 (=) Abs2 Rep2"  | 
595  | 
shows "Quotient3 (R1 OOO (=)) (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: 
47094 
diff
changeset
 | 
596  | 
using assms  | 
| 47308 | 597  | 
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: 
47094 
diff
changeset
 | 
598  | 
|
| 60758 | 599  | 
subsection \<open>Quotient3 to Quotient\<close>  | 
| 
47362
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
600  | 
|
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
601  | 
lemma Quotient3_to_Quotient:  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
602  | 
assumes "Quotient3 R Abs Rep"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
603  | 
and "T \<equiv> \<lambda>x y. R x x \<and> Abs x = y"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
604  | 
shows "Quotient R Abs Rep T"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
605  | 
using assms unfolding Quotient3_def by (intro QuotientI) blast+  | 
| 
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: 
47094 
diff
changeset
 | 
606  | 
|
| 
47362
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
607  | 
lemma Quotient3_to_Quotient_equivp:  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
608  | 
assumes q: "Quotient3 R Abs Rep"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
609  | 
and T_def: "T \<equiv> \<lambda>x y. Abs x = y"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
610  | 
and eR: "equivp R"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
611  | 
shows "Quotient R Abs Rep T"  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
612  | 
proof (intro QuotientI)  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
613  | 
fix a  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
614  | 
show "Abs (Rep a) = a" using q by(rule Quotient3_abs_rep)  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
615  | 
next  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
616  | 
fix a  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
617  | 
show "R (Rep a) (Rep a)" using q by(rule Quotient3_rep_reflp)  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
618  | 
next  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
619  | 
fix r s  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
620  | 
show "R r s = (R r r \<and> R s s \<and> Abs r = Abs s)" using q by(rule Quotient3_rel[symmetric])  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
621  | 
next  | 
| 
 
b1f099bdfbba
connect the Quotient package to the Lifting package
 
kuncar 
parents: 
47361 
diff
changeset
 | 
622  | 
show "T = (\<lambda>x y. R x x \<and> Abs x = y)" using T_def equivp_reflp[OF eR] by simp  | 
| 
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: 
47094 
diff
changeset
 | 
623  | 
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: 
47094 
diff
changeset
 | 
624  | 
|
| 60758 | 625  | 
subsection \<open>ML setup\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
626  | 
|
| 60758 | 627  | 
text \<open>Auxiliary data for the quotient package\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
628  | 
|
| 57960 | 629  | 
named_theorems quot_equiv "equivalence relation theorems"  | 
| 59028 | 630  | 
and quot_respect "respectfulness theorems"  | 
631  | 
and quot_preserve "preservation theorems"  | 
|
632  | 
and id_simps "identity simp rules for maps"  | 
|
633  | 
and quot_thm "quotient theorems"  | 
|
| 69605 | 634  | 
ML_file \<open>Tools/Quotient/quotient_info.ML\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
635  | 
|
| 55945 | 636  | 
declare [[mapQ3 "fun" = (rel_fun, fun_quotient3)]]  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
637  | 
|
| 47308 | 638  | 
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: 
44413 
diff
changeset
 | 
639  | 
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: 
44413 
diff
changeset
 | 
640  | 
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
 | 
641  | 
lemmas [quot_equiv] = identity_equivp  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
642  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
643  | 
|
| 60758 | 644  | 
text \<open>Lemmas about simplifying id's.\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
645  | 
lemmas [id_simps] =  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
646  | 
id_def[symmetric]  | 
| 40602 | 647  | 
map_fun_id  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
648  | 
id_apply  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
649  | 
id_o  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
650  | 
o_id  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
651  | 
eq_comp_r  | 
| 
44413
 
80d460bc6fa8
Quotient Package: some infrastructure for lifting inside sets
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
44242 
diff
changeset
 | 
652  | 
vimage_id  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
653  | 
|
| 60758 | 654  | 
text \<open>Translation functions for the lifting process.\<close>  | 
| 69605 | 655  | 
ML_file \<open>Tools/Quotient/quotient_term.ML\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
656  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
657  | 
|
| 60758 | 658  | 
text \<open>Definitions of the quotient types.\<close>  | 
| 69605 | 659  | 
ML_file \<open>Tools/Quotient/quotient_type.ML\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
660  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
661  | 
|
| 60758 | 662  | 
text \<open>Definitions for quotient constants.\<close>  | 
| 69605 | 663  | 
ML_file \<open>Tools/Quotient/quotient_def.ML\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
664  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
665  | 
|
| 60758 | 666  | 
text \<open>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
667  | 
An auxiliary constant for recording some information  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
668  | 
about the lifted theorem in a tactic.  | 
| 60758 | 669  | 
\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
670  | 
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: 
40031 
diff
changeset
 | 
671  | 
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: 
40031 
diff
changeset
 | 
672  | 
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: 
40031 
diff
changeset
 | 
673  | 
"Quot_True x \<longleftrightarrow> True"  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
674  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
675  | 
lemma  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
676  | 
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
 | 
677  | 
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
 | 
678  | 
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
 | 
679  | 
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
 | 
680  | 
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
 | 
681  | 
by (simp_all add: Quot_True_def ext)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
682  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
683  | 
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
 | 
684  | 
by (simp add: Quot_True_def)  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
685  | 
|
| 63343 | 686  | 
context includes lifting_syntax  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
687  | 
begin  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
688  | 
|
| 60758 | 689  | 
text \<open>Tactics for proving the lifted theorems\<close>  | 
| 69605 | 690  | 
ML_file \<open>Tools/Quotient/quotient_tacs.ML\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
691  | 
|
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
692  | 
end  | 
| 
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
693  | 
|
| 60758 | 694  | 
subsection \<open>Methods / Interface\<close>  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
695  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
696  | 
method_setup lifting =  | 
| 60758 | 697  | 
\<open>Attrib.thms >> (fn thms => fn ctxt =>  | 
698  | 
SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms))\<close>  | 
|
699  | 
\<open>lift theorems to quotient types\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
700  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
701  | 
method_setup lifting_setup =  | 
| 60758 | 702  | 
\<open>Attrib.thm >> (fn thm => fn ctxt =>  | 
703  | 
SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm))\<close>  | 
|
704  | 
\<open>set up the three goals for the quotient lifting procedure\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
705  | 
|
| 
37593
 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
37564 
diff
changeset
 | 
706  | 
method_setup descending =  | 
| 60758 | 707  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt []))\<close>  | 
708  | 
\<open>decend theorems to the raw level\<close>  | 
|
| 
37593
 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
37564 
diff
changeset
 | 
709  | 
|
| 
 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
37564 
diff
changeset
 | 
710  | 
method_setup descending_setup =  | 
| 60758 | 711  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt []))\<close>  | 
712  | 
\<open>set up the three goals for the decending theorems\<close>  | 
|
| 
37593
 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
37564 
diff
changeset
 | 
713  | 
|
| 
45782
 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
45680 
diff
changeset
 | 
714  | 
method_setup partiality_descending =  | 
| 60758 | 715  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt []))\<close>  | 
716  | 
\<open>decend theorems to the raw level\<close>  | 
|
| 
45782
 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
45680 
diff
changeset
 | 
717  | 
|
| 
 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
45680 
diff
changeset
 | 
718  | 
method_setup partiality_descending_setup =  | 
| 60758 | 719  | 
\<open>Scan.succeed (fn ctxt =>  | 
720  | 
SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt []))\<close>  | 
|
721  | 
\<open>set up the three goals for the decending theorems\<close>  | 
|
| 
45782
 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
45680 
diff
changeset
 | 
722  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
723  | 
method_setup regularize =  | 
| 60758 | 724  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt))\<close>  | 
725  | 
\<open>prove the regularization goals from the quotient lifting procedure\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
726  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
727  | 
method_setup injection =  | 
| 60758 | 728  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt))\<close>  | 
729  | 
\<open>prove the rep/abs injection goals from the quotient lifting procedure\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
730  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
731  | 
method_setup cleaning =  | 
| 60758 | 732  | 
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt))\<close>  | 
733  | 
\<open>prove the cleaning goals from the quotient lifting procedure\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
734  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
735  | 
attribute_setup quot_lifted =  | 
| 60758 | 736  | 
\<open>Scan.succeed Quotient_Tacs.lifted_attrib\<close>  | 
737  | 
\<open>lift theorems to quotient types\<close>  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
738  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
739  | 
no_notation  | 
| 
53011
 
aeee0a4be6cf
introduce locale with syntax for fun_rel and map_fun and make thus ===> and ---> local
 
kuncar 
parents: 
51112 
diff
changeset
 | 
740  | 
rel_conj (infixr "OOO" 75)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
741  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
742  | 
end  | 
| 
47488
 
be6dd389639d
centralized enriched_type declaration, thanks to in-situ available Isar commands
 
haftmann 
parents: 
47436 
diff
changeset
 | 
743  |