src/HOL/Lifting.thy
author kuncar
Thu, 10 Apr 2014 17:48:18 +0200
changeset 57866 f4ba736040fa
parent 57861 c1048f5bbb45
child 58740 882091eb1e9a
permissions -rw-r--r--
setup for Transfer and Lifting from BNF; tuned thm names
     1 (*  Title:      HOL/Lifting.thy
     2     Author:     Brian Huffman and Ondrej Kuncar
     3     Author:     Cezary Kaliszyk and Christian Urban
     4 *)
     5 
     6 header {* Lifting package *}
     7 
     8 theory Lifting
     9 imports Equiv_Relations Transfer
    10 keywords
    11   "parametric" and
    12   "print_quot_maps" "print_quotients" :: diag and
    13   "lift_definition" :: thy_goal and
    14   "setup_lifting" "lifting_forget" "lifting_update" :: thy_decl
    15 begin
    16 
    17 subsection {* Function map *}
    18 
    19 context
    20 begin
    21 interpretation lifting_syntax .
    22 
    23 lemma map_fun_id:
    24   "(id ---> id) = id"
    25   by (simp add: fun_eq_iff)
    26 
    27 subsection {* Quotient Predicate *}
    28 
    29 definition
    30   "Quotient R Abs Rep T \<longleftrightarrow>
    31      (\<forall>a. Abs (Rep a) = a) \<and> 
    32      (\<forall>a. R (Rep a) (Rep a)) \<and>
    33      (\<forall>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s) \<and>
    34      T = (\<lambda>x y. R x x \<and> Abs x = y)"
    35 
    36 lemma QuotientI:
    37   assumes "\<And>a. Abs (Rep a) = a"
    38     and "\<And>a. R (Rep a) (Rep a)"
    39     and "\<And>r s. R r s \<longleftrightarrow> R r r \<and> R s s \<and> Abs r = Abs s"
    40     and "T = (\<lambda>x y. R x x \<and> Abs x = y)"
    41   shows "Quotient R Abs Rep T"
    42   using assms unfolding Quotient_def by blast
    43 
    44 context
    45   fixes R Abs Rep T
    46   assumes a: "Quotient R Abs Rep T"
    47 begin
    48 
    49 lemma Quotient_abs_rep: "Abs (Rep a) = a"
    50   using a unfolding Quotient_def
    51   by simp
    52 
    53 lemma Quotient_rep_reflp: "R (Rep a) (Rep a)"
    54   using a unfolding Quotient_def
    55   by blast
    56 
    57 lemma Quotient_rel:
    58   "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *}
    59   using a unfolding Quotient_def
    60   by blast
    61 
    62 lemma Quotient_cr_rel: "T = (\<lambda>x y. R x x \<and> Abs x = y)"
    63   using a unfolding Quotient_def
    64   by blast
    65 
    66 lemma Quotient_refl1: "R r s \<Longrightarrow> R r r"
    67   using a unfolding Quotient_def
    68   by fast
    69 
    70 lemma Quotient_refl2: "R r s \<Longrightarrow> R s s"
    71   using a unfolding Quotient_def
    72   by fast
    73 
    74 lemma Quotient_rel_rep: "R (Rep a) (Rep b) \<longleftrightarrow> a = b"
    75   using a unfolding Quotient_def
    76   by metis
    77 
    78 lemma Quotient_rep_abs: "R r r \<Longrightarrow> R (Rep (Abs r)) r"
    79   using a unfolding Quotient_def
    80   by blast
    81 
    82 lemma Quotient_rep_abs_eq: "R t t \<Longrightarrow> R \<le> op= \<Longrightarrow> Rep (Abs t) = t"
    83   using a unfolding Quotient_def
    84   by blast
    85 
    86 lemma Quotient_rep_abs_fold_unmap: 
    87   assumes "x' \<equiv> Abs x" and "R x x" and "Rep x' \<equiv> Rep' x'" 
    88   shows "R (Rep' x') x"
    89 proof -
    90   have "R (Rep x') x" using assms(1-2) Quotient_rep_abs by auto
    91   then show ?thesis using assms(3) by simp
    92 qed
    93 
    94 lemma Quotient_Rep_eq:
    95   assumes "x' \<equiv> Abs x" 
    96   shows "Rep x' \<equiv> Rep x'"
    97 by simp
    98 
    99 lemma Quotient_rel_abs: "R r s \<Longrightarrow> Abs r = Abs s"
   100   using a unfolding Quotient_def
   101   by blast
   102 
   103 lemma Quotient_rel_abs2:
   104   assumes "R (Rep x) y"
   105   shows "x = Abs y"
   106 proof -
   107   from assms have "Abs (Rep x) = Abs y" by (auto intro: Quotient_rel_abs)
   108   then show ?thesis using assms(1) by (simp add: Quotient_abs_rep)
   109 qed
   110 
   111 lemma Quotient_symp: "symp R"
   112   using a unfolding Quotient_def using sympI by (metis (full_types))
   113 
   114 lemma Quotient_transp: "transp R"
   115   using a unfolding Quotient_def using transpI by (metis (full_types))
   116 
   117 lemma Quotient_part_equivp: "part_equivp R"
   118 by (metis Quotient_rep_reflp Quotient_symp Quotient_transp part_equivpI)
   119 
   120 end
   121 
   122 lemma identity_quotient: "Quotient (op =) id id (op =)"
   123 unfolding Quotient_def by simp 
   124 
   125 text {* TODO: Use one of these alternatives as the real definition. *}
   126 
   127 lemma Quotient_alt_def:
   128   "Quotient R Abs Rep T \<longleftrightarrow>
   129     (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>
   130     (\<forall>b. T (Rep b) b) \<and>
   131     (\<forall>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)"
   132 apply safe
   133 apply (simp (no_asm_use) only: Quotient_def, fast)
   134 apply (simp (no_asm_use) only: Quotient_def, fast)
   135 apply (simp (no_asm_use) only: Quotient_def, fast)
   136 apply (simp (no_asm_use) only: Quotient_def, fast)
   137 apply (simp (no_asm_use) only: Quotient_def, fast)
   138 apply (simp (no_asm_use) only: Quotient_def, fast)
   139 apply (rule QuotientI)
   140 apply simp
   141 apply metis
   142 apply simp
   143 apply (rule ext, rule ext, metis)
   144 done
   145 
   146 lemma Quotient_alt_def2:
   147   "Quotient R Abs Rep T \<longleftrightarrow>
   148     (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and>
   149     (\<forall>b. T (Rep b) b) \<and>
   150     (\<forall>x y. R x y \<longleftrightarrow> T x (Abs y) \<and> T y (Abs x))"
   151   unfolding Quotient_alt_def by (safe, metis+)
   152 
   153 lemma Quotient_alt_def3:
   154   "Quotient R Abs Rep T \<longleftrightarrow>
   155     (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and>
   156     (\<forall>x y. R x y \<longleftrightarrow> (\<exists>z. T x z \<and> T y z))"
   157   unfolding Quotient_alt_def2 by (safe, metis+)
   158 
   159 lemma Quotient_alt_def4:
   160   "Quotient R Abs Rep T \<longleftrightarrow>
   161     (\<forall>a b. T a b \<longrightarrow> Abs a = b) \<and> (\<forall>b. T (Rep b) b) \<and> R = T OO conversep T"
   162   unfolding Quotient_alt_def3 fun_eq_iff by auto
   163 
   164 lemma Quotient_alt_def5:
   165   "Quotient R Abs Rep T \<longleftrightarrow>
   166     T \<le> BNF_Util.Grp UNIV Abs \<and> BNF_Util.Grp UNIV Rep \<le> T\<inverse>\<inverse> \<and> R = T OO T\<inverse>\<inverse>"
   167   unfolding Quotient_alt_def4 Grp_def by blast
   168 
   169 lemma fun_quotient:
   170   assumes 1: "Quotient R1 abs1 rep1 T1"
   171   assumes 2: "Quotient R2 abs2 rep2 T2"
   172   shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2) (T1 ===> T2)"
   173   using assms unfolding Quotient_alt_def2
   174   unfolding rel_fun_def fun_eq_iff map_fun_apply
   175   by (safe, metis+)
   176 
   177 lemma apply_rsp:
   178   fixes f g::"'a \<Rightarrow> 'c"
   179   assumes q: "Quotient R1 Abs1 Rep1 T1"
   180   and     a: "(R1 ===> R2) f g" "R1 x y"
   181   shows "R2 (f x) (g y)"
   182   using a by (auto elim: rel_funE)
   183 
   184 lemma apply_rsp':
   185   assumes a: "(R1 ===> R2) f g" "R1 x y"
   186   shows "R2 (f x) (g y)"
   187   using a by (auto elim: rel_funE)
   188 
   189 lemma apply_rsp'':
   190   assumes "Quotient R Abs Rep T"
   191   and "(R ===> S) f f"
   192   shows "S (f (Rep x)) (f (Rep x))"
   193 proof -
   194   from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient_rep_reflp)
   195   then show ?thesis using assms(2) by (auto intro: apply_rsp')
   196 qed
   197 
   198 subsection {* Quotient composition *}
   199 
   200 lemma Quotient_compose:
   201   assumes 1: "Quotient R1 Abs1 Rep1 T1"
   202   assumes 2: "Quotient R2 Abs2 Rep2 T2"
   203   shows "Quotient (T1 OO R2 OO conversep T1) (Abs2 \<circ> Abs1) (Rep1 \<circ> Rep2) (T1 OO T2)"
   204   using assms unfolding Quotient_alt_def4 by fastforce
   205 
   206 lemma equivp_reflp2:
   207   "equivp R \<Longrightarrow> reflp R"
   208   by (erule equivpE)
   209 
   210 subsection {* Respects predicate *}
   211 
   212 definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
   213   where "Respects R = {x. R x x}"
   214 
   215 lemma in_respects: "x \<in> Respects R \<longleftrightarrow> R x x"
   216   unfolding Respects_def by simp
   217 
   218 lemma UNIV_typedef_to_Quotient:
   219   assumes "type_definition Rep Abs UNIV"
   220   and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
   221   shows "Quotient (op =) Abs Rep T"
   222 proof -
   223   interpret type_definition Rep Abs UNIV by fact
   224   from Abs_inject Rep_inverse Abs_inverse T_def show ?thesis 
   225     by (fastforce intro!: QuotientI fun_eq_iff)
   226 qed
   227 
   228 lemma UNIV_typedef_to_equivp:
   229   fixes Abs :: "'a \<Rightarrow> 'b"
   230   and Rep :: "'b \<Rightarrow> 'a"
   231   assumes "type_definition Rep Abs (UNIV::'a set)"
   232   shows "equivp (op=::'a\<Rightarrow>'a\<Rightarrow>bool)"
   233 by (rule identity_equivp)
   234 
   235 lemma typedef_to_Quotient:
   236   assumes "type_definition Rep Abs S"
   237   and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
   238   shows "Quotient (eq_onp (\<lambda>x. x \<in> S)) Abs Rep T"
   239 proof -
   240   interpret type_definition Rep Abs S by fact
   241   from Rep Abs_inject Rep_inverse Abs_inverse T_def show ?thesis
   242     by (auto intro!: QuotientI simp: eq_onp_def fun_eq_iff)
   243 qed
   244 
   245 lemma typedef_to_part_equivp:
   246   assumes "type_definition Rep Abs S"
   247   shows "part_equivp (eq_onp (\<lambda>x. x \<in> S))"
   248 proof (intro part_equivpI)
   249   interpret type_definition Rep Abs S by fact
   250   show "\<exists>x. eq_onp (\<lambda>x. x \<in> S) x x" using Rep by (auto simp: eq_onp_def)
   251 next
   252   show "symp (eq_onp (\<lambda>x. x \<in> S))" by (auto intro: sympI simp: eq_onp_def)
   253 next
   254   show "transp (eq_onp (\<lambda>x. x \<in> S))" by (auto intro: transpI simp: eq_onp_def)
   255 qed
   256 
   257 lemma open_typedef_to_Quotient:
   258   assumes "type_definition Rep Abs {x. P x}"
   259   and T_def: "T \<equiv> (\<lambda>x y. x = Rep y)"
   260   shows "Quotient (eq_onp P) Abs Rep T"
   261   using typedef_to_Quotient [OF assms] by simp
   262 
   263 lemma open_typedef_to_part_equivp:
   264   assumes "type_definition Rep Abs {x. P x}"
   265   shows "part_equivp (eq_onp P)"
   266   using typedef_to_part_equivp [OF assms] by simp
   267 
   268 text {* Generating transfer rules for quotients. *}
   269 
   270 context
   271   fixes R Abs Rep T
   272   assumes 1: "Quotient R Abs Rep T"
   273 begin
   274 
   275 lemma Quotient_right_unique: "right_unique T"
   276   using 1 unfolding Quotient_alt_def right_unique_def by metis
   277 
   278 lemma Quotient_right_total: "right_total T"
   279   using 1 unfolding Quotient_alt_def right_total_def by metis
   280 
   281 lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)"
   282   using 1 unfolding Quotient_alt_def rel_fun_def by simp
   283 
   284 lemma Quotient_abs_induct:
   285   assumes "\<And>y. R y y \<Longrightarrow> P (Abs y)" shows "P x"
   286   using 1 assms unfolding Quotient_def by metis
   287 
   288 end
   289 
   290 text {* Generating transfer rules for total quotients. *}
   291 
   292 context
   293   fixes R Abs Rep T
   294   assumes 1: "Quotient R Abs Rep T" and 2: "reflp R"
   295 begin
   296 
   297 lemma Quotient_left_total: "left_total T"
   298   using 1 2 unfolding Quotient_alt_def left_total_def reflp_def by auto
   299 
   300 lemma Quotient_bi_total: "bi_total T"
   301   using 1 2 unfolding Quotient_alt_def bi_total_def reflp_def by auto
   302 
   303 lemma Quotient_id_abs_transfer: "(op = ===> T) (\<lambda>x. x) Abs"
   304   using 1 2 unfolding Quotient_alt_def reflp_def rel_fun_def by simp
   305 
   306 lemma Quotient_total_abs_induct: "(\<And>y. P (Abs y)) \<Longrightarrow> P x"
   307   using 1 2 assms unfolding Quotient_alt_def reflp_def by metis
   308 
   309 lemma Quotient_total_abs_eq_iff: "Abs x = Abs y \<longleftrightarrow> R x y"
   310   using Quotient_rel [OF 1] 2 unfolding reflp_def by simp
   311 
   312 end
   313 
   314 text {* Generating transfer rules for a type defined with @{text "typedef"}. *}
   315 
   316 context
   317   fixes Rep Abs A T
   318   assumes type: "type_definition Rep Abs A"
   319   assumes T_def: "T \<equiv> (\<lambda>(x::'a) (y::'b). x = Rep y)"
   320 begin
   321 
   322 lemma typedef_left_unique: "left_unique T"
   323   unfolding left_unique_def T_def
   324   by (simp add: type_definition.Rep_inject [OF type])
   325 
   326 lemma typedef_bi_unique: "bi_unique T"
   327   unfolding bi_unique_def T_def
   328   by (simp add: type_definition.Rep_inject [OF type])
   329 
   330 (* the following two theorems are here only for convinience *)
   331 
   332 lemma typedef_right_unique: "right_unique T"
   333   using T_def type Quotient_right_unique typedef_to_Quotient 
   334   by blast
   335 
   336 lemma typedef_right_total: "right_total T"
   337   using T_def type Quotient_right_total typedef_to_Quotient 
   338   by blast
   339 
   340 lemma typedef_rep_transfer: "(T ===> op =) (\<lambda>x. x) Rep"
   341   unfolding rel_fun_def T_def by simp
   342 
   343 end
   344 
   345 text {* Generating the correspondence rule for a constant defined with
   346   @{text "lift_definition"}. *}
   347 
   348 lemma Quotient_to_transfer:
   349   assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c"
   350   shows "T c c'"
   351   using assms by (auto dest: Quotient_cr_rel)
   352 
   353 text {* Proving reflexivity *}
   354 
   355 lemma Quotient_to_left_total:
   356   assumes q: "Quotient R Abs Rep T"
   357   and r_R: "reflp R"
   358   shows "left_total T"
   359 using r_R Quotient_cr_rel[OF q] unfolding left_total_def by (auto elim: reflpE)
   360 
   361 lemma Quotient_composition_ge_eq:
   362   assumes "left_total T"
   363   assumes "R \<ge> op="
   364   shows "(T OO R OO T\<inverse>\<inverse>) \<ge> op="
   365 using assms unfolding left_total_def by fast
   366 
   367 lemma Quotient_composition_le_eq:
   368   assumes "left_unique T"
   369   assumes "R \<le> op="
   370   shows "(T OO R OO T\<inverse>\<inverse>) \<le> op="
   371 using assms unfolding left_unique_def by blast
   372 
   373 lemma eq_onp_le_eq:
   374   "eq_onp P \<le> op=" unfolding eq_onp_def by blast
   375 
   376 lemma reflp_ge_eq:
   377   "reflp R \<Longrightarrow> R \<ge> op=" unfolding reflp_def by blast
   378 
   379 lemma ge_eq_refl:
   380   "R \<ge> op= \<Longrightarrow> R x x" by blast
   381 
   382 text {* Proving a parametrized correspondence relation *}
   383 
   384 definition POS :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   385 "POS A B \<equiv> A \<le> B"
   386 
   387 definition  NEG :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
   388 "NEG A B \<equiv> B \<le> A"
   389 
   390 lemma pos_OO_eq:
   391   shows "POS (A OO op=) A"
   392 unfolding POS_def OO_def by blast
   393 
   394 lemma pos_eq_OO:
   395   shows "POS (op= OO A) A"
   396 unfolding POS_def OO_def by blast
   397 
   398 lemma neg_OO_eq:
   399   shows "NEG (A OO op=) A"
   400 unfolding NEG_def OO_def by auto
   401 
   402 lemma neg_eq_OO:
   403   shows "NEG (op= OO A) A"
   404 unfolding NEG_def OO_def by blast
   405 
   406 lemma POS_trans:
   407   assumes "POS A B"
   408   assumes "POS B C"
   409   shows "POS A C"
   410 using assms unfolding POS_def by auto
   411 
   412 lemma NEG_trans:
   413   assumes "NEG A B"
   414   assumes "NEG B C"
   415   shows "NEG A C"
   416 using assms unfolding NEG_def by auto
   417 
   418 lemma POS_NEG:
   419   "POS A B \<equiv> NEG B A"
   420   unfolding POS_def NEG_def by auto
   421 
   422 lemma NEG_POS:
   423   "NEG A B \<equiv> POS B A"
   424   unfolding POS_def NEG_def by auto
   425 
   426 lemma POS_pcr_rule:
   427   assumes "POS (A OO B) C"
   428   shows "POS (A OO B OO X) (C OO X)"
   429 using assms unfolding POS_def OO_def by blast
   430 
   431 lemma NEG_pcr_rule:
   432   assumes "NEG (A OO B) C"
   433   shows "NEG (A OO B OO X) (C OO X)"
   434 using assms unfolding NEG_def OO_def by blast
   435 
   436 lemma POS_apply:
   437   assumes "POS R R'"
   438   assumes "R f g"
   439   shows "R' f g"
   440 using assms unfolding POS_def by auto
   441 
   442 text {* Proving a parametrized correspondence relation *}
   443 
   444 lemma fun_mono:
   445   assumes "A \<ge> C"
   446   assumes "B \<le> D"
   447   shows   "(A ===> B) \<le> (C ===> D)"
   448 using assms unfolding rel_fun_def by blast
   449 
   450 lemma pos_fun_distr: "((R ===> S) OO (R' ===> S')) \<le> ((R OO R') ===> (S OO S'))"
   451 unfolding OO_def rel_fun_def by blast
   452 
   453 lemma functional_relation: "right_unique R \<Longrightarrow> left_total R \<Longrightarrow> \<forall>x. \<exists>!y. R x y"
   454 unfolding right_unique_def left_total_def by blast
   455 
   456 lemma functional_converse_relation: "left_unique R \<Longrightarrow> right_total R \<Longrightarrow> \<forall>y. \<exists>!x. R x y"
   457 unfolding left_unique_def right_total_def by blast
   458 
   459 lemma neg_fun_distr1:
   460 assumes 1: "left_unique R" "right_total R"
   461 assumes 2: "right_unique R'" "left_total R'"
   462 shows "(R OO R' ===> S OO S') \<le> ((R ===> S) OO (R' ===> S')) "
   463   using functional_relation[OF 2] functional_converse_relation[OF 1]
   464   unfolding rel_fun_def OO_def
   465   apply clarify
   466   apply (subst all_comm)
   467   apply (subst all_conj_distrib[symmetric])
   468   apply (intro choice)
   469   by metis
   470 
   471 lemma neg_fun_distr2:
   472 assumes 1: "right_unique R'" "left_total R'"
   473 assumes 2: "left_unique S'" "right_total S'"
   474 shows "(R OO R' ===> S OO S') \<le> ((R ===> S) OO (R' ===> S'))"
   475   using functional_converse_relation[OF 2] functional_relation[OF 1]
   476   unfolding rel_fun_def OO_def
   477   apply clarify
   478   apply (subst all_comm)
   479   apply (subst all_conj_distrib[symmetric])
   480   apply (intro choice)
   481   by metis
   482 
   483 subsection {* Domains *}
   484 
   485 lemma composed_equiv_rel_eq_onp:
   486   assumes "left_unique R"
   487   assumes "(R ===> op=) P P'"
   488   assumes "Domainp R = P''"
   489   shows "(R OO eq_onp P' OO R\<inverse>\<inverse>) = eq_onp (inf P'' P)"
   490 using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def rel_fun_def eq_onp_def
   491 fun_eq_iff by blast
   492 
   493 lemma composed_equiv_rel_eq_eq_onp:
   494   assumes "left_unique R"
   495   assumes "Domainp R = P"
   496   shows "(R OO op= OO R\<inverse>\<inverse>) = eq_onp P"
   497 using assms unfolding OO_def conversep_iff Domainp_iff[abs_def] left_unique_def eq_onp_def
   498 fun_eq_iff is_equality_def by metis
   499 
   500 lemma pcr_Domainp_par_left_total:
   501   assumes "Domainp B = P"
   502   assumes "left_total A"
   503   assumes "(A ===> op=) P' P"
   504   shows "Domainp (A OO B) = P'"
   505 using assms
   506 unfolding Domainp_iff[abs_def] OO_def bi_unique_def left_total_def rel_fun_def 
   507 by (fast intro: fun_eq_iff)
   508 
   509 lemma pcr_Domainp_par:
   510 assumes "Domainp B = P2"
   511 assumes "Domainp A = P1"
   512 assumes "(A ===> op=) P2' P2"
   513 shows "Domainp (A OO B) = (inf P1 P2')"
   514 using assms unfolding rel_fun_def Domainp_iff[abs_def] OO_def
   515 by (fast intro: fun_eq_iff)
   516 
   517 definition rel_pred_comp :: "('a => 'b => bool) => ('b => bool) => 'a => bool"
   518 where "rel_pred_comp R P \<equiv> \<lambda>x. \<exists>y. R x y \<and> P y"
   519 
   520 lemma pcr_Domainp:
   521 assumes "Domainp B = P"
   522 shows "Domainp (A OO B) = (\<lambda>x. \<exists>y. A x y \<and> P y)"
   523 using assms by blast
   524 
   525 lemma pcr_Domainp_total:
   526   assumes "left_total B"
   527   assumes "Domainp A = P"
   528   shows "Domainp (A OO B) = P"
   529 using assms unfolding left_total_def 
   530 by fast
   531 
   532 lemma Quotient_to_Domainp:
   533   assumes "Quotient R Abs Rep T"
   534   shows "Domainp T = (\<lambda>x. R x x)"  
   535 by (simp add: Domainp_iff[abs_def] Quotient_cr_rel[OF assms])
   536 
   537 lemma eq_onp_to_Domainp:
   538   assumes "Quotient (eq_onp P) Abs Rep T"
   539   shows "Domainp T = P"
   540 by (simp add: eq_onp_def Domainp_iff[abs_def] Quotient_cr_rel[OF assms])
   541 
   542 end
   543 
   544 subsection {* ML setup *}
   545 
   546 ML_file "Tools/Lifting/lifting_util.ML"
   547 
   548 ML_file "Tools/Lifting/lifting_info.ML"
   549 setup Lifting_Info.setup
   550 
   551 (* setup for the function type *)
   552 declare fun_quotient[quot_map]
   553 declare fun_mono[relator_mono]
   554 lemmas [relator_distr] = pos_fun_distr neg_fun_distr1 neg_fun_distr2
   555 
   556 ML_file "Tools/Lifting/lifting_bnf.ML"
   557 
   558 ML_file "Tools/Lifting/lifting_term.ML"
   559 
   560 ML_file "Tools/Lifting/lifting_def.ML"
   561 
   562 ML_file "Tools/Lifting/lifting_setup.ML"
   563                            
   564 hide_const (open) POS NEG
   565 
   566 end