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