src/HOL/Lifting.thy
author haftmann
Mon, 28 May 2012 13:38:07 +0200
changeset 48003 1d11af40b106
parent 47982 7aa35601ff65
child 48891 c0eafbd55de3
permissions -rw-r--r--
dropped sort constraints on datatype specifications
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Lifting.thy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     2
    Author:     Brian Huffman and Ondrej Kuncar
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     3
    Author:     Cezary Kaliszyk and Christian Urban
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     4
*)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     5
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     6
header {* Lifting package *}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     7
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     8
theory Lifting
47325
ec6187036495 new transfer proof method
huffman
parents: 47308
diff changeset
     9
imports Plain Equiv_Relations Transfer
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    10
keywords
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    11
  "print_quotmaps" "print_quotients" :: diag and
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    12
  "lift_definition" :: thy_goal and
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    13
  "setup_lifting" :: thy_decl
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    14
uses
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47652
diff changeset
    15
  ("Tools/Lifting/lifting_util.ML")
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    16
  ("Tools/Lifting/lifting_info.ML")
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    17
  ("Tools/Lifting/lifting_term.ML")
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    18
  ("Tools/Lifting/lifting_def.ML")
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    19
  ("Tools/Lifting/lifting_setup.ML")
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    20
begin
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    21
47325
ec6187036495 new transfer proof method
huffman
parents: 47308
diff changeset
    22
subsection {* Function map *}
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    23
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    24
notation map_fun (infixr "--->" 55)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    25
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    26
lemma map_fun_id:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    27
  "(id ---> id) = id"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    28
  by (simp add: fun_eq_iff)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    29
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    30
subsection {* Quotient Predicate *}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    31
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    32
definition
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    33
  "Quotient R Abs Rep T \<longleftrightarrow>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    34
     (\<forall>a. Abs (Rep a) = a) \<and> 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    35
     (\<forall>a. R (Rep a) (Rep a)) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    36
     (\<forall>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    37
     T = (\<lambda>x y. R x x \<and> Abs x = y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    38
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    39
lemma QuotientI:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    40
  assumes "\<And>a. Abs (Rep a) = a"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    41
    and "\<And>a. R (Rep a) (Rep a)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    42
    and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    43
    and "T = (\<lambda>x y. R x x \<and> Abs x = y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    44
  shows "Quotient R Abs Rep T"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    45
  using assms unfolding Quotient_def by blast
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    46
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    47
context
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    48
  fixes R Abs Rep T
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    49
  assumes a: "Quotient R Abs Rep T"
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    50
begin
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    51
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    52
lemma Quotient_abs_rep: "Abs (Rep a) = a"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    53
  using a unfolding Quotient_def
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    54
  by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    55
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    56
lemma Quotient_rep_reflp: "R (Rep a) (Rep a)"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    57
  using a unfolding Quotient_def
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    58
  by blast
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    59
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    60
lemma Quotient_rel:
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    61
  "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *}
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    62
  using a unfolding Quotient_def
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    63
  by blast
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    64
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    65
lemma Quotient_cr_rel: "T = (\<lambda>x y. R x x \<and> Abs x = y)"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    66
  using a unfolding Quotient_def
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    67
  by blast
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    68
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    69
lemma Quotient_refl1: "R r s \<Longrightarrow> R r r"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    70
  using a unfolding Quotient_def
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    71
  by fast
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    72
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    73
lemma Quotient_refl2: "R r s \<Longrightarrow> R s s"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    74
  using a unfolding Quotient_def
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    75
  by fast
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    76
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    77
lemma Quotient_rel_rep: "R (Rep a) (Rep b) \<longleftrightarrow> a = b"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    78
  using a unfolding Quotient_def
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    79
  by metis
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    80
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    81
lemma Quotient_rep_abs: "R r r \<Longrightarrow> R (Rep (Abs r)) r"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    82
  using a unfolding Quotient_def
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    83
  by blast
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    84
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    85
lemma Quotient_rep_abs_fold_unmap: 
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    86
  assumes "x' \<equiv> Abs x" and "R x x" and "Rep x' \<equiv> Rep' x'" 
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    87
  shows "R (Rep' x') x"
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    88
proof -
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    89
  have "R (Rep x') x" using assms(1-2) Quotient_rep_abs by auto
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    90
  then show ?thesis using assms(3) by simp
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    91
qed
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    92
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    93
lemma Quotient_Rep_eq:
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    94
  assumes "x' \<equiv> Abs x" 
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    95
  shows "Rep x' \<equiv> Rep x'"
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    96
by simp
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
    97
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    98
lemma Quotient_rel_abs: "R r s \<Longrightarrow> Abs r = Abs s"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
    99
  using a unfolding Quotient_def
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   100
  by blast
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   101
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   102
lemma Quotient_rel_abs2:
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   103
  assumes "R (Rep x) y"
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   104
  shows "x = Abs y"
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   105
proof -
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   106
  from assms have "Abs (Rep x) = Abs y" by (auto intro: Quotient_rel_abs)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   107
  then show ?thesis using assms(1) by (simp add: Quotient_abs_rep)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   108
qed
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47936
diff changeset
   109
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   110
lemma Quotient_symp: "symp R"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   111
  using a unfolding Quotient_def using sympI by (metis (full_types))
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   112
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   113
lemma Quotient_transp: "transp R"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   114
  using a unfolding Quotient_def using transpI by (metis (full_types))
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   115
47536
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   116
lemma Quotient_part_equivp: "part_equivp R"
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   117
by (metis Quotient_rep_reflp Quotient_symp Quotient_transp part_equivpI)
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   118
8474a865a4e5 use context block
huffman
parents: 47535
diff changeset
   119
end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   120
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   121
lemma identity_quotient: "Quotient (op =) id id (op =)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   122
unfolding Quotient_def by simp 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   123
47652
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   124
text {* TODO: Use one of these alternatives as the real definition. *}
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   125
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   126
lemma Quotient_alt_def:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   127
  "Quotient R Abs Rep T \<longleftrightarrow>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   128
    (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   129
    (\<forall>b. T (Rep b) b) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   130
    (\<forall>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   131
apply safe
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   132
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   133
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   134
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   135
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   136
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   137
apply (simp (no_asm_use) only: Quotient_def, fast)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   138
apply (rule QuotientI)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   139
apply simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   140
apply metis
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   141
apply simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   142
apply (rule ext, rule ext, metis)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   143
done
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   144
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   145
lemma Quotient_alt_def2:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   146
  "Quotient R Abs Rep T \<longleftrightarrow>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   147
    (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   148
    (\<forall>b. T (Rep b) b) \<and>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   149
    (\<forall>x y. R x y \<longleftrightarrow> T x (Abs y) \<and> T y (Abs x))"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   150
  unfolding Quotient_alt_def by (safe, metis+)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   151
47652
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   152
lemma Quotient_alt_def3:
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   153
  "Quotient R Abs Rep T \<longleftrightarrow>
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   154
    (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and>
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   155
    (\<forall>x y. R x y \<longleftrightarrow> (\<exists>z. T x z \<and> T y z))"
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   156
  unfolding Quotient_alt_def2 by (safe, metis+)
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   157
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   158
lemma Quotient_alt_def4:
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   159
  "Quotient R Abs Rep T \<longleftrightarrow>
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   160
    (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and> R = T OO conversep T"
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   161
  unfolding Quotient_alt_def3 fun_eq_iff by auto
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   162
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   163
lemma fun_quotient:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   164
  assumes 1: "Quotient R1 abs1 rep1 T1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   165
  assumes 2: "Quotient R2 abs2 rep2 T2"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   166
  shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2) (T1 ===> T2)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   167
  using assms unfolding Quotient_alt_def2
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   168
  unfolding fun_rel_def fun_eq_iff map_fun_apply
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   169
  by (safe, metis+)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   170
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   171
lemma apply_rsp:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   172
  fixes f g::"'a \<Rightarrow> 'c"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   173
  assumes q: "Quotient R1 Abs1 Rep1 T1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   174
  and     a: "(R1 ===> R2) f g" "R1 x y"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   175
  shows "R2 (f x) (g y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   176
  using a by (auto elim: fun_relE)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   177
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   178
lemma apply_rsp':
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   179
  assumes a: "(R1 ===> R2) f g" "R1 x y"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   180
  shows "R2 (f x) (g y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   181
  using a by (auto elim: fun_relE)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   182
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   183
lemma apply_rsp'':
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   184
  assumes "Quotient R Abs Rep T"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   185
  and "(R ===> S) f f"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   186
  shows "S (f (Rep x)) (f (Rep x))"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   187
proof -
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   188
  from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   189
  then show ?thesis using assms(2) by (auto intro: apply_rsp')
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   190
qed
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   191
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   192
subsection {* Quotient composition *}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   193
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   194
lemma Quotient_compose:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   195
  assumes 1: "Quotient R1 Abs1 Rep1 T1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   196
  assumes 2: "Quotient R2 Abs2 Rep2 T2"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   197
  shows "Quotient (T1 OO R2 OO conversep T1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2) (T1 OO T2)"
47652
1b722b100301 move alternative definition lemmas into Lifting.thy;
huffman
parents: 47651
diff changeset
   198
  using assms unfolding Quotient_alt_def4 by (auto intro!: ext)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   199
47521
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47501
diff changeset
   200
lemma equivp_reflp2:
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47501
diff changeset
   201
  "equivp R \<Longrightarrow> reflp R"
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47501
diff changeset
   202
  by (erule equivpE)
69f95ac85c3d tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents: 47501
diff changeset
   203
47544
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   204
subsection {* Respects predicate *}
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   205
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   206
definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   207
  where "Respects R = {x. R x x}"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   208
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   209
lemma in_respects: "x \<in> Respects R \<longleftrightarrow> R x x"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   210
  unfolding Respects_def by simp
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   211
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   212
subsection {* Invariant *}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   213
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   214
definition invariant :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   215
  where "invariant R = (\<lambda>x y. R x \<and> x = y)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   216
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   217
lemma invariant_to_eq:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   218
  assumes "invariant P x y"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   219
  shows "x = y"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   220
using assms by (simp add: invariant_def)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   221
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   222
lemma fun_rel_eq_invariant:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   223
  shows "((invariant R) ===> S) = (\<lambda>f g. \<forall>x. R x \<longrightarrow> S (f x) (g x))"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   224
by (auto simp add: invariant_def fun_rel_def)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   225
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   226
lemma invariant_same_args:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   227
  shows "invariant P x x \<equiv> P x"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   228
using assms by (auto simp add: invariant_def)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   229
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   230
lemma UNIV_typedef_to_Quotient:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   231
  assumes "type_definition Rep Abs UNIV"
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   232
  and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   233
  shows "Quotient (op =) Abs Rep T"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   234
proof -
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   235
  interpret type_definition Rep Abs UNIV by fact
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   236
  from Abs_inject Rep_inverse Abs_inverse T_def show ?thesis 
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   237
    by (fastforce intro!: QuotientI fun_eq_iff)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   238
qed
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   239
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   240
lemma UNIV_typedef_to_equivp:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   241
  fixes Abs :: "'a \<Rightarrow> 'b"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   242
  and Rep :: "'b \<Rightarrow> 'a"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   243
  assumes "type_definition Rep Abs (UNIV::'a set)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   244
  shows "equivp (op=::'a\<Rightarrow>'a\<Rightarrow>bool)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   245
by (rule identity_equivp)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   246
47354
95846613e414 update keywords file
huffman
parents: 47351
diff changeset
   247
lemma typedef_to_Quotient:
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   248
  assumes "type_definition Rep Abs S"
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   249
  and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
47501
0b9294e093db leave Lifting prefix
kuncar
parents: 47436
diff changeset
   250
  shows "Quotient (invariant (\<lambda>x. x \<in> S)) Abs Rep T"
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   251
proof -
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   252
  interpret type_definition Rep Abs S by fact
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   253
  from Rep Abs_inject Rep_inverse Abs_inverse T_def show ?thesis
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   254
    by (auto intro!: QuotientI simp: invariant_def fun_eq_iff)
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   255
qed
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   256
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   257
lemma typedef_to_part_equivp:
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   258
  assumes "type_definition Rep Abs S"
47501
0b9294e093db leave Lifting prefix
kuncar
parents: 47436
diff changeset
   259
  shows "part_equivp (invariant (\<lambda>x. x \<in> S))"
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   260
proof (intro part_equivpI)
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   261
  interpret type_definition Rep Abs S by fact
47501
0b9294e093db leave Lifting prefix
kuncar
parents: 47436
diff changeset
   262
  show "\<exists>x. invariant (\<lambda>x. x \<in> S) x x" using Rep by (auto simp: invariant_def)
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   263
next
47501
0b9294e093db leave Lifting prefix
kuncar
parents: 47436
diff changeset
   264
  show "symp (invariant (\<lambda>x. x \<in> S))" by (auto intro: sympI simp: invariant_def)
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   265
next
47501
0b9294e093db leave Lifting prefix
kuncar
parents: 47436
diff changeset
   266
  show "transp (invariant (\<lambda>x. x \<in> S))" by (auto intro: transpI simp: invariant_def)
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   267
qed
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   268
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   269
lemma open_typedef_to_Quotient:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   270
  assumes "type_definition Rep Abs {x. P x}"
47354
95846613e414 update keywords file
huffman
parents: 47351
diff changeset
   271
  and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   272
  shows "Quotient (invariant P) Abs Rep T"
47651
8e4f50afd21a tuned proofs
huffman
parents: 47575
diff changeset
   273
  using typedef_to_Quotient [OF assms] by simp
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   274
47361
87c0eaf04bad support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents: 47354
diff changeset
   275
lemma open_typedef_to_part_equivp:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   276
  assumes "type_definition Rep Abs {x. P x}"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   277
  shows "part_equivp (invariant P)"
47651
8e4f50afd21a tuned proofs
huffman
parents: 47575
diff changeset
   278
  using typedef_to_part_equivp [OF assms] by simp
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   279
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   280
text {* Generating transfer rules for quotients. *}
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   281
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   282
context
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   283
  fixes R Abs Rep T
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   284
  assumes 1: "Quotient R Abs Rep T"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   285
begin
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   286
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   287
lemma Quotient_right_unique: "right_unique T"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   288
  using 1 unfolding Quotient_alt_def right_unique_def by metis
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   289
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   290
lemma Quotient_right_total: "right_total T"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   291
  using 1 unfolding Quotient_alt_def right_total_def by metis
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   292
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   293
lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   294
  using 1 unfolding Quotient_alt_def fun_rel_def by simp
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   295
47538
1f0ec5b8135a add lemma Quotient_abs_induct
huffman
parents: 47537
diff changeset
   296
lemma Quotient_abs_induct:
1f0ec5b8135a add lemma Quotient_abs_induct
huffman
parents: 47537
diff changeset
   297
  assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x"
1f0ec5b8135a add lemma Quotient_abs_induct
huffman
parents: 47537
diff changeset
   298
  using 1 assms unfolding Quotient_def by metis
1f0ec5b8135a add lemma Quotient_abs_induct
huffman
parents: 47537
diff changeset
   299
47544
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   300
lemma Quotient_All_transfer:
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   301
  "((T ===> op =) ===> op =) (Ball (Respects R)) All"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   302
  unfolding fun_rel_def Respects_def Quotient_cr_rel [OF 1]
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   303
  by (auto, metis Quotient_abs_induct)
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   304
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   305
lemma Quotient_Ex_transfer:
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   306
  "((T ===> op =) ===> op =) (Bex (Respects R)) Ex"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   307
  unfolding fun_rel_def Respects_def Quotient_cr_rel [OF 1]
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   308
  by (auto, metis Quotient_rep_reflp [OF 1] Quotient_abs_rep [OF 1])
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   309
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   310
lemma Quotient_forall_transfer:
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   311
  "((T ===> op =) ===> op =) (transfer_bforall (\<lambda>x. R x x)) transfer_forall"
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   312
  using Quotient_All_transfer
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   313
  unfolding transfer_forall_def transfer_bforall_def
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   314
    Ball_def [abs_def] in_respects .
e455cdaac479 move constant 'Respects' into Lifting.thy;
huffman
parents: 47538
diff changeset
   315
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   316
end
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   317
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   318
text {* Generating transfer rules for total quotients. *}
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   319
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   320
context
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   321
  fixes R Abs Rep T
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   322
  assumes 1: "Quotient R Abs Rep T" and 2: "reflp R"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   323
begin
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   324
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   325
lemma Quotient_bi_total: "bi_total T"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   326
  using 1 2 unfolding Quotient_alt_def bi_total_def reflp_def by auto
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   327
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   328
lemma Quotient_id_abs_transfer: "(op = ===> T) (\<lambda>x. x) Abs"
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   329
  using 1 2 unfolding Quotient_alt_def reflp_def fun_rel_def by simp
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   330
47575
b90cd7016d4f generate abs_induct rules for quotient types
huffman
parents: 47545
diff changeset
   331
lemma Quotient_total_abs_induct: "(\<And>y. P (Abs y)) \<Longrightarrow> P x"
b90cd7016d4f generate abs_induct rules for quotient types
huffman
parents: 47545
diff changeset
   332
  using 1 2 assms unfolding Quotient_alt_def reflp_def by metis
b90cd7016d4f generate abs_induct rules for quotient types
huffman
parents: 47545
diff changeset
   333
47889
29212a4bb866 lifting package produces abs_eq_iff rules for total quotients
huffman
parents: 47777
diff changeset
   334
lemma Quotient_total_abs_eq_iff: "Abs x = Abs y \<longleftrightarrow> R x y"
29212a4bb866 lifting package produces abs_eq_iff rules for total quotients
huffman
parents: 47777
diff changeset
   335
  using Quotient_rel [OF 1] 2 unfolding reflp_def by simp
29212a4bb866 lifting package produces abs_eq_iff rules for total quotients
huffman
parents: 47777
diff changeset
   336
47537
b06be48923a4 more usage of context blocks
huffman
parents: 47536
diff changeset
   337
end
47376
776254f89a18 add transfer lemmas for quotients
huffman
parents: 47369
diff changeset
   338
47368
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   339
text {* Generating transfer rules for a type defined with @{text "typedef"}. *}
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   340
47534
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   341
context
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   342
  fixes Rep Abs A T
47368
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   343
  assumes type: "type_definition Rep Abs A"
47534
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   344
  assumes T_def: "T \<equiv> (\<lambda>(x::'a) (y::'b). x = Rep y)"
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   345
begin
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   346
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   347
lemma typedef_bi_unique: "bi_unique T"
47368
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   348
  unfolding bi_unique_def T_def
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   349
  by (simp add: type_definition.Rep_inject [OF type])
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   350
47535
0f94b02fda1c lifting_setup generates transfer rule for rep of typedefs
huffman
parents: 47534
diff changeset
   351
lemma typedef_rep_transfer: "(T ===> op =) (\<lambda>x. x) Rep"
0f94b02fda1c lifting_setup generates transfer rule for rep of typedefs
huffman
parents: 47534
diff changeset
   352
  unfolding fun_rel_def T_def by simp
0f94b02fda1c lifting_setup generates transfer rule for rep of typedefs
huffman
parents: 47534
diff changeset
   353
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   354
lemma typedef_All_transfer: "((T ===> op =) ===> op =) (Ball A) All"
47534
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   355
  unfolding T_def fun_rel_def
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   356
  by (metis type_definition.Rep [OF type]
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   357
    type_definition.Abs_inverse [OF type])
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   358
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   359
lemma typedef_Ex_transfer: "((T ===> op =) ===> op =) (Bex A) Ex"
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   360
  unfolding T_def fun_rel_def
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   361
  by (metis type_definition.Rep [OF type]
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   362
    type_definition.Abs_inverse [OF type])
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   363
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   364
lemma typedef_forall_transfer:
47534
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   365
  "((T ===> op =) ===> op =)
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   366
    (transfer_bforall (\<lambda>x. x \<in> A)) transfer_forall"
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   367
  unfolding transfer_bforall_def transfer_forall_def Ball_def [symmetric]
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47544
diff changeset
   368
  by (rule typedef_All_transfer)
47534
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   369
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   370
end
06cc372a80ed use context block to organize typedef lifting theorems
huffman
parents: 47521
diff changeset
   371
47368
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   372
text {* Generating the correspondence rule for a constant defined with
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   373
  @{text "lift_definition"}. *}
4c522dff1f4d add lemmas for generating transfer rules for typedefs
huffman
parents: 47354
diff changeset
   374
47351
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47325
diff changeset
   375
lemma Quotient_to_transfer:
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47325
diff changeset
   376
  assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c"
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47325
diff changeset
   377
  shows "T c c'"
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47325
diff changeset
   378
  using assms by (auto dest: Quotient_cr_rel)
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47325
diff changeset
   379
47982
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   380
text {* Proving reflexivity *}
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   381
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   382
definition left_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   383
  where "left_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y)"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   384
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   385
lemma left_totalI:
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   386
  "(\<And>x. \<exists>y. R x y) \<Longrightarrow> left_total R"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   387
unfolding left_total_def by blast
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   388
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   389
lemma left_totalE:
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   390
  assumes "left_total R"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   391
  obtains "(\<And>x. \<exists>y. R x y)"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   392
using assms unfolding left_total_def by blast
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   393
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   394
lemma Quotient_to_left_total:
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   395
  assumes q: "Quotient R Abs Rep T"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   396
  and r_R: "reflp R"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   397
  shows "left_total T"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   398
using r_R Quotient_cr_rel[OF q] unfolding left_total_def by (auto elim: reflpE)
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   399
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   400
lemma reflp_Quotient_composition:
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   401
  assumes lt_R1: "left_total R1"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   402
  and r_R2: "reflp R2"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   403
  shows "reflp (R1 OO R2 OO R1\<inverse>\<inverse>)"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   404
using assms
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   405
proof -
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   406
  {
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   407
    fix x
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   408
    from lt_R1 obtain y where "R1 x y" unfolding left_total_def by blast
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   409
    moreover then have "R1\<inverse>\<inverse> y x" by simp
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   410
    moreover have "R2 y y" using r_R2 by (auto elim: reflpE)
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   411
    ultimately have "(R1 OO R2 OO R1\<inverse>\<inverse>) x x" by auto
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   412
  }
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   413
  then show ?thesis by (auto intro: reflpI)
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   414
qed
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   415
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   416
lemma reflp_equality: "reflp (op =)"
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   417
by (auto intro: reflpI)
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   418
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   419
subsection {* ML setup *}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   420
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47652
diff changeset
   421
use "Tools/Lifting/lifting_util.ML"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   422
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   423
use "Tools/Lifting/lifting_info.ML"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   424
setup Lifting_Info.setup
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   425
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47698
diff changeset
   426
declare fun_quotient[quot_map]
47982
7aa35601ff65 prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents: 47937
diff changeset
   427
lemmas [reflexivity_rule] = reflp_equality reflp_Quotient_composition
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   428
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   429
use "Tools/Lifting/lifting_term.ML"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   430
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   431
use "Tools/Lifting/lifting_def.ML"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   432
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   433
use "Tools/Lifting/lifting_setup.ML"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   434
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   435
hide_const (open) invariant
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   436
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   437
end