QuotScript.thy
changeset 542 fe468f8723fc
parent 541 94deffed619d
child 543 d030c8e19465
equal deleted inserted replaced
541:94deffed619d 542:fe468f8723fc
    70   assumes a: "Quotient R Abs Rep"
    70   assumes a: "Quotient R Abs Rep"
    71   shows "R r r \<Longrightarrow> R (Rep (Abs r)) r"
    71   shows "R r r \<Longrightarrow> R (Rep (Abs r)) r"
    72 using a unfolding Quotient_def
    72 using a unfolding Quotient_def
    73 by blast
    73 by blast
    74 
    74 
    75 lemma IDENTITY_equivp:
    75 lemma identity_equivp:
    76   shows "equivp (op =)"
    76   shows "equivp (op =)"
    77 unfolding equivp_def
    77 unfolding equivp_def
    78 by auto
    78 by auto
    79 
    79 
    80 lemma IDENTITY_Quotient:
    80 lemma identity_quotient:
    81   shows "Quotient (op =) id id"
    81   shows "Quotient (op =) id id"
    82 unfolding Quotient_def id_def
    82 unfolding Quotient_def id_def
    83 by blast
    83 by blast
    84 
    84 
    85 lemma Quotient_symp:
    85 lemma Quotient_symp:
   111   "f ---> g \<equiv> fun_map f g"
   111   "f ---> g \<equiv> fun_map f g"
   112 
   112 
   113 lemma fun_map_id:
   113 lemma fun_map_id:
   114   shows "(id ---> id) = id"
   114   shows "(id ---> id) = id"
   115 by (simp add: expand_fun_eq id_def)
   115 by (simp add: expand_fun_eq id_def)
   116 
       
   117 (* Not used *)
       
   118 lemma in_fun:
       
   119   shows "x \<in> ((f ---> g) s) = g (f x \<in> s)"
       
   120 by (simp add: mem_def)
       
   121 
   116 
   122 fun
   117 fun
   123   fun_rel
   118   fun_rel
   124 where
   119 where
   125   "fun_rel E1 E2 f g = (\<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"
   120   "fun_rel E1 E2 f g = (\<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"
   170 definition
   165 definition
   171   Respects
   166   Respects
   172 where
   167 where
   173   "Respects R x \<equiv> (R x x)"
   168   "Respects R x \<equiv> (R x x)"
   174 
   169 
   175 lemma IN_RESPECTS:
   170 lemma in_respects:
   176   shows "(x \<in> Respects R) = R x x"
   171   shows "(x \<in> Respects R) = R x x"
   177 unfolding mem_def Respects_def by simp
   172 unfolding mem_def Respects_def by simp
       
   173 
       
   174 (* TODO: it is the same as APPLY_RSP *)
       
   175 (* q1 and q2 not used; see next lemma *)
       
   176 lemma fun_rel_MP:
       
   177   assumes q1: "Quotient R1 Abs1 Rep1"
       
   178   and     q2: "Quotient R2 Abs2 Rep2"
       
   179   shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
       
   180 by simp
       
   181 
       
   182 lemma fun_rel_IMP:
       
   183   shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
       
   184 by simp
       
   185 
       
   186 
       
   187 lemma equals_rsp:
       
   188   assumes q: "Quotient R Abs Rep"
       
   189   and     a: "R xa xb" "R ya yb"
       
   190   shows "R xa ya = R xb yb"
       
   191 using Quotient_symp[OF q] Quotient_transp[OF q] unfolding symp_def transp_def
       
   192 using a by blast
       
   193 
       
   194 lemma lambda_prs:
       
   195   assumes q1: "Quotient R1 Abs1 Rep1"
       
   196   and     q2: "Quotient R2 Abs2 Rep2"
       
   197   shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)"
       
   198 unfolding expand_fun_eq
       
   199 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
       
   200 by simp
       
   201 
       
   202 lemma lambda_prs1:
       
   203   assumes q1: "Quotient R1 Abs1 Rep1"
       
   204   and     q2: "Quotient R2 Abs2 Rep2"
       
   205   shows "(Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x) = (\<lambda>x. f x)"
       
   206 unfolding expand_fun_eq
       
   207 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
       
   208 by simp
       
   209 
       
   210 lemma rep_abs_rsp:
       
   211   assumes q: "Quotient R Abs Rep"
       
   212   and     a: "R x1 x2"
       
   213   shows "R x1 (Rep (Abs x2))"
       
   214 using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
       
   215 
       
   216 (* ----------------------------------------------------- *)
       
   217 (* Quantifiers: FORALL, EXISTS, EXISTS_UNIQUE,           *)
       
   218 (*              Ball, Bex, RES_EXISTS_EQUIV              *)
       
   219 (* ----------------------------------------------------- *)
       
   220 
       
   221 (* bool theory: COND, LET *)
       
   222 
       
   223 lemma IF_PRS:
       
   224   assumes q: "Quotient R Abs Rep"
       
   225   shows "If a b c = Abs (If a (Rep b) (Rep c))"
       
   226 using Quotient_abs_rep[OF q] by auto
       
   227 
       
   228 (* ask peter: no use of q *)
       
   229 lemma IF_RSP:
       
   230   assumes q: "Quotient R Abs Rep"
       
   231   and     a: "a1 = a2" "R b1 b2" "R c1 c2"
       
   232   shows "R (If a1 b1 c1) (If a2 b2 c2)"
       
   233 using a by auto
       
   234 
       
   235 lemma LET_PRS:
       
   236   assumes q1: "Quotient R1 Abs1 Rep1"
       
   237   and     q2: "Quotient R2 Abs2 Rep2"
       
   238   shows "Let x f = Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f))"
       
   239 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
       
   240 
       
   241 lemma LET_RSP:
       
   242   assumes q1: "Quotient R1 Abs1 Rep1"
       
   243   and     q2: "Quotient R2 Abs2 Rep2"
       
   244   and     a1: "(R1 ===> R2) f g"
       
   245   and     a2: "R1 x y"
       
   246   shows "R2 (Let x f) (Let y g)"
       
   247 using fun_rel_MP[OF q1 q2 a1] a2
       
   248 by auto
       
   249 
       
   250 
       
   251 (* ask peter what are literal_case *)
       
   252 (* literal_case_PRS *)
       
   253 (* literal_case_RSP *)
       
   254 
       
   255 
       
   256 (* FUNCTION APPLICATION *)
       
   257 
       
   258 (* In the following theorem R1 can be instantiated with anything,
       
   259    but we know some of the types of the Rep and Abs functions;
       
   260    so by solving Quotient assumptions we can get a unique R2 that
       
   261    will be provable; which is why we need to use APPLY_RSP *)
       
   262 lemma apply_rsp:
       
   263   assumes q: "Quotient R1 Abs1 Rep1"
       
   264   and     a: "(R1 ===> R2) f g" "R1 x y"
       
   265   shows "R2 ((f::'a\<Rightarrow>'c) x) ((g::'a\<Rightarrow>'c) y)"
       
   266 using a by (rule fun_rel_IMP)
       
   267 
       
   268 lemma apply_rsp':
       
   269   assumes a: "(R1 ===> R2) f g" "R1 x y"
       
   270   shows "R2 (f x) (g y)"
       
   271 using a by (rule fun_rel_IMP)
       
   272 
       
   273 
       
   274 (* combinators: I, K, o, C, W *)
       
   275 
       
   276 (* We use id_simps which includes id_apply; so these 2 theorems can be removed *)
       
   277 lemma I_PRS:
       
   278   assumes q: "Quotient R Abs Rep"
       
   279   shows "id e = Abs (id (Rep e))"
       
   280 using Quotient_abs_rep[OF q] by auto
       
   281 
       
   282 lemma I_RSP:
       
   283   assumes q: "Quotient R Abs Rep"
       
   284   and     a: "R e1 e2"
       
   285   shows "R (id e1) (id e2)"
       
   286 using a by auto
       
   287 
       
   288 lemma o_PRS:
       
   289   assumes q1: "Quotient R1 Abs1 Rep1"
       
   290   and     q2: "Quotient R2 Abs2 Rep2"
       
   291   and     q3: "Quotient R3 Abs3 Rep3"
       
   292   shows "f o g = (Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g))"
       
   293 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
       
   294 unfolding o_def expand_fun_eq
       
   295 by simp
       
   296 
       
   297 lemma o_RSP:
       
   298   assumes q1: "Quotient R1 Abs1 Rep1"
       
   299   and     q2: "Quotient R2 Abs2 Rep2"
       
   300   and     q3: "Quotient R3 Abs3 Rep3"
       
   301   and     a1: "(R2 ===> R3) f1 f2"
       
   302   and     a2: "(R1 ===> R2) g1 g2"
       
   303   shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
       
   304 using a1 a2 unfolding o_def expand_fun_eq
       
   305 by (auto)
       
   306 
       
   307 
       
   308 
       
   309 
       
   310 
       
   311 lemma COND_PRS:
       
   312   assumes a: "Quotient R absf repf"
       
   313   shows "(if a then b else c) = absf (if a then repf b else repf c)"
       
   314   using a unfolding Quotient_def by auto
       
   315 
       
   316 
       
   317 
       
   318 
       
   319 
       
   320 (* Set of lemmas for regularisation of ball and bex *)
       
   321 lemma ball_reg_eqv:
       
   322   fixes P :: "'a \<Rightarrow> bool"
       
   323   assumes a: "equivp R"
       
   324   shows "Ball (Respects R) P = (All P)"
       
   325   by (metis equivp_def in_respects a)
       
   326 
       
   327 lemma bex_reg_eqv:
       
   328   fixes P :: "'a \<Rightarrow> bool"
       
   329   assumes a: "equivp R"
       
   330   shows "Bex (Respects R) P = (Ex P)"
       
   331   by (metis equivp_def in_respects a)
       
   332 
       
   333 lemma ball_reg_right:
       
   334   assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x"
       
   335   shows "All P \<longrightarrow> Ball R Q"
       
   336   by (metis COMBC_def Collect_def Collect_mem_eq a)
       
   337 
       
   338 lemma bex_reg_left:
       
   339   assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x"
       
   340   shows "Bex R Q \<longrightarrow> Ex P"
       
   341   by (metis COMBC_def Collect_def Collect_mem_eq a)
       
   342 
       
   343 lemma ball_reg_left:
       
   344   assumes a: "equivp R"
       
   345   shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P"
       
   346   by (metis equivp_reflp in_respects a)
       
   347 
       
   348 lemma bex_reg_right:
       
   349   assumes a: "equivp R"
       
   350   shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P"
       
   351   by (metis equivp_reflp in_respects a)
       
   352 
       
   353 lemma ball_reg_eqv_range:
       
   354   fixes P::"'a \<Rightarrow> bool"
       
   355   and x::"'a"
       
   356   assumes a: "equivp R2"
       
   357   shows   "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
       
   358   apply(rule iffI)
       
   359   apply(rule allI)
       
   360   apply(drule_tac x="\<lambda>y. f x" in bspec)
       
   361   apply(simp add: Respects_def in_respects)
       
   362   apply(rule impI)
       
   363   using a equivp_reflp_symp_transp[of "R2"]
       
   364   apply(simp add: reflp_def)
       
   365   apply(simp)
       
   366   apply(simp)
       
   367   done
       
   368 
       
   369 lemma bex_reg_eqv_range:
       
   370   fixes P::"'a \<Rightarrow> bool"
       
   371   and x::"'a"
       
   372   assumes a: "equivp R2"
       
   373   shows   "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))"
       
   374   apply(auto)
       
   375   apply(rule_tac x="\<lambda>y. f x" in bexI)
       
   376   apply(simp)
       
   377   apply(simp add: Respects_def in_respects)
       
   378   apply(rule impI)
       
   379   using a equivp_reflp_symp_transp[of "R2"]
       
   380   apply(simp add: reflp_def)
       
   381   done
       
   382 
       
   383 lemma all_reg:
       
   384   assumes a: "!x :: 'a. (P x --> Q x)"
       
   385   and     b: "All P"
       
   386   shows "All Q"
       
   387   using a b by (metis)
       
   388 
       
   389 lemma ex_reg:
       
   390   assumes a: "!x :: 'a. (P x --> Q x)"
       
   391   and     b: "Ex P"
       
   392   shows "Ex Q"
       
   393   using a b by (metis)
       
   394 
       
   395 lemma ball_reg:
       
   396   assumes a: "!x :: 'a. (R x --> P x --> Q x)"
       
   397   and     b: "Ball R P"
       
   398   shows "Ball R Q"
       
   399   using a b by (metis COMBC_def Collect_def Collect_mem_eq)
       
   400 
       
   401 lemma bex_reg:
       
   402   assumes a: "!x :: 'a. (R x --> P x --> Q x)"
       
   403   and     b: "Bex R P"
       
   404   shows "Bex R Q"
       
   405   using a b by (metis COMBC_def Collect_def Collect_mem_eq)
       
   406 
       
   407 lemma ball_all_comm:
       
   408   "(\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)) \<Longrightarrow> ((\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y))"
       
   409 by auto
       
   410 
       
   411 lemma bex_ex_comm:
       
   412   "((\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)) \<Longrightarrow> ((\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y))"
       
   413 by auto
       
   414 
       
   415 (* 2 lemmas needed for proving repabs_inj *)
       
   416 lemma ball_rsp:
       
   417   assumes a: "(R ===> (op =)) f g"
       
   418   shows "Ball (Respects R) f = Ball (Respects R) g"
       
   419   using a by (simp add: Ball_def in_respects)
       
   420 
       
   421 lemma bex_rsp:
       
   422   assumes a: "(R ===> (op =)) f g"
       
   423   shows "(Bex (Respects R) f = Bex (Respects R) g)"
       
   424   using a by (simp add: Bex_def in_respects)
       
   425 
       
   426 (* 2 lemmas needed for cleaning of quantifiers *)
       
   427 lemma all_prs:
       
   428   assumes a: "Quotient R absf repf"
       
   429   shows "Ball (Respects R) ((absf ---> id) f) = All f"
       
   430   using a unfolding Quotient_def
       
   431   by (metis in_respects fun_map.simps id_apply)
       
   432 
       
   433 lemma ex_prs:
       
   434   assumes a: "Quotient R absf repf"
       
   435   shows "Bex (Respects R) ((absf ---> id) f) = Ex f"
       
   436   using a unfolding Quotient_def
       
   437   by (metis COMBC_def Collect_def Collect_mem_eq in_respects fun_map.simps id_apply)
       
   438 
       
   439 
       
   440 (* UNUSED *)
       
   441 lemma Quotient_rel_abs:
       
   442   assumes a: "Quotient E Abs Rep"
       
   443   shows "E r s \<Longrightarrow> Abs r = Abs s"
       
   444 using a unfolding Quotient_def
       
   445 by blast
       
   446 
       
   447 lemma Quotient_rel_abs_eq:
       
   448   assumes a: "Quotient E Abs Rep"
       
   449   shows "E r r \<Longrightarrow> E s s \<Longrightarrow> E r s = (Abs r = Abs s)"
       
   450 using a unfolding Quotient_def
       
   451 by blast
       
   452 
       
   453 lemma in_fun:
       
   454   shows "x \<in> ((f ---> g) s) = g (f x \<in> s)"
       
   455 by (simp add: mem_def)
   178 
   456 
   179 lemma RESPECTS_THM:
   457 lemma RESPECTS_THM:
   180   shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))"
   458   shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))"
   181 unfolding Respects_def
   459 unfolding Respects_def
   182 by (simp add: expand_fun_eq) 
   460 by (simp add: expand_fun_eq) 
   183 
       
   184 lemma RESPECTS_MP:
       
   185   assumes a: "Respects (R1 ===> R2) f"
       
   186   and     b: "R1 x y"
       
   187   shows "R2 (f x) (f y)"
       
   188 using a b unfolding Respects_def
       
   189 by simp
       
   190 
   461 
   191 lemma RESPECTS_REP_ABS:
   462 lemma RESPECTS_REP_ABS:
   192   assumes a: "Quotient R1 Abs1 Rep1"
   463   assumes a: "Quotient R1 Abs1 Rep1"
   193   and     b: "Respects (R1 ===> R2) f"
   464   and     b: "Respects (R1 ===> R2) f"
   194   and     c: "R1 x x"
   465   and     c: "R1 x x"
   195   shows "R2 (f (Rep1 (Abs1 x))) (f x)"
   466   shows "R2 (f (Rep1 (Abs1 x))) (f x)"
   196 using a b[simplified RESPECTS_THM] c unfolding Quotient_def
   467 using a b[simplified RESPECTS_THM] c unfolding Quotient_def
   197 by blast
   468 by blast
   198 
   469 
       
   470 lemma RESPECTS_MP:
       
   471   assumes a: "Respects (R1 ===> R2) f"
       
   472   and     b: "R1 x y"
       
   473   shows "R2 (f x) (f y)"
       
   474 using a b unfolding Respects_def
       
   475 by simp
       
   476 
   199 lemma RESPECTS_o:
   477 lemma RESPECTS_o:
   200   assumes a: "Respects (R2 ===> R3) f"
   478   assumes a: "Respects (R2 ===> R3) f"
   201   and     b: "Respects (R1 ===> R2) g"
   479   and     b: "Respects (R1 ===> R2) g"
   202   shows "Respects (R1 ===> R3) (f o g)"
   480   shows "Respects (R1 ===> R3) (f o g)"
   203 using a b unfolding Respects_def
   481 using a b unfolding Respects_def
   204 by simp
   482 by simp
   205 
   483 
   206 (*
       
   207 definition
       
   208   "RES_EXISTS_EQUIV R P \<equiv> (\<exists>x \<in> Respects R. P x) \<and> 
       
   209                           (\<forall>x\<in> Respects R. \<forall>y\<in> Respects R. P x \<and> P y \<longrightarrow> R x y)"
       
   210 *)
       
   211 
       
   212 lemma fun_rel_EQ_REL:
   484 lemma fun_rel_EQ_REL:
   213   assumes q1: "Quotient R1 Abs1 Rep1"
   485   assumes q1: "Quotient R1 Abs1 Rep1"
   214   and     q2: "Quotient R2 Abs2 Rep2"
   486   and     q2: "Quotient R2 Abs2 Rep2"
   215   shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) 
   487   shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) 
   216                              \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))"
   488                              \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))"
   217 using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq
   489 using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq
   218 by blast
   490 by blast
   219 
   491 
   220 (* TODO: it is the same as APPLY_RSP *)
   492 (* Not used since in the end we just unfold fun_map *)
   221 (* q1 and q2 not used; see next lemma *)
   493 lemma APP_PRS:
   222 lemma fun_rel_MP:
   494   assumes q1: "Quotient R1 abs1 rep1"
   223   assumes q1: "Quotient R1 Abs1 Rep1"
   495   and     q2: "Quotient R2 abs2 rep2"
   224   and     q2: "Quotient R2 Abs2 Rep2"
   496   shows "abs2 ((abs1 ---> rep2) f (rep1 x)) = f x"
   225   shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
   497 unfolding expand_fun_eq
   226 by simp
   498 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
   227 
   499 by simp
   228 lemma fun_rel_IMP:
   500 
   229   shows "(R1 ===> R2) f g \<Longrightarrow> R1 x y \<Longrightarrow> R2 (f x) (g y)"
   501 (* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *)
   230 by simp
   502 lemma LAMBDA_RSP:
       
   503   assumes q1: "Quotient R1 Abs1 Rep1"
       
   504   and     q2: "Quotient R2 Abs2 Rep2"
       
   505   and     a: "(R1 ===> R2) f1 f2"
       
   506   shows "(R1 ===> R2) (\<lambda>x. f1 x) (\<lambda>y. f2 y)"
       
   507 by (rule a)
       
   508 
       
   509 (* ASK Peter about next four lemmas in quotientScript
       
   510 lemma ABSTRACT_PRS:
       
   511   assumes q1: "Quotient R1 Abs1 Rep1"
       
   512   and     q2: "Quotient R2 Abs2 Rep2"
       
   513   shows "f = (Rep1 ---> Abs2) ???"
       
   514 *)
       
   515 
   231 
   516 
   232 lemma fun_rel_EQUALS:
   517 lemma fun_rel_EQUALS:
   233   assumes q1: "Quotient R1 Abs1 Rep1"
   518   assumes q1: "Quotient R1 Abs1 Rep1"
   234   and     q2: "Quotient R2 Abs2 Rep2"
   519   and     q2: "Quotient R2 Abs2 Rep2"
   235   and     r1: "Respects (R1 ===> R2) f"
   520   and     r1: "Respects (R1 ===> R2) f"
   252   and     a:  "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g"
   537   and     a:  "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g"
   253   shows "R1 x y \<Longrightarrow> R2 (f x) (g y)"
   538   shows "R1 x y \<Longrightarrow> R2 (f x) (g y)"
   254 using q1 q2 r1 r2 a
   539 using q1 q2 r1 r2 a
   255 by (simp add: fun_rel_EQUALS)
   540 by (simp add: fun_rel_EQUALS)
   256 
   541 
   257 lemma equals_rsp:
       
   258   assumes q: "Quotient R Abs Rep"
       
   259   and     a: "R xa xb" "R ya yb"
       
   260   shows "R xa ya = R xb yb"
       
   261 using Quotient_symp[OF q] Quotient_transp[OF q] unfolding symp_def transp_def
       
   262 using a by blast
       
   263 
       
   264 lemma lambda_prs:
       
   265   assumes q1: "Quotient R1 Abs1 Rep1"
       
   266   and     q2: "Quotient R2 Abs2 Rep2"
       
   267   shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)"
       
   268 unfolding expand_fun_eq
       
   269 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
       
   270 by simp
       
   271 
       
   272 lemma lambda_prs1:
       
   273   assumes q1: "Quotient R1 Abs1 Rep1"
       
   274   and     q2: "Quotient R2 Abs2 Rep2"
       
   275   shows "(Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x) = (\<lambda>x. f x)"
       
   276 unfolding expand_fun_eq
       
   277 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
       
   278 by simp
       
   279 
       
   280 (* Not used since applic_prs proves a version for an arbitrary number of arguments *)
       
   281 lemma APP_PRS:
       
   282   assumes q1: "Quotient R1 abs1 rep1"
       
   283   and     q2: "Quotient R2 abs2 rep2"
       
   284   shows "abs2 ((abs1 ---> rep2) f (rep1 x)) = f x"
       
   285 unfolding expand_fun_eq
       
   286 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
       
   287 by simp
       
   288 
       
   289 (* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *)
       
   290 lemma LAMBDA_RSP:
       
   291   assumes q1: "Quotient R1 Abs1 Rep1"
       
   292   and     q2: "Quotient R2 Abs2 Rep2"
       
   293   and     a: "(R1 ===> R2) f1 f2"
       
   294   shows "(R1 ===> R2) (\<lambda>x. f1 x) (\<lambda>y. f2 y)"
       
   295 by (rule a)
       
   296 
       
   297 (* ASK Peter about next four lemmas in quotientScript
       
   298 lemma ABSTRACT_PRS:
       
   299   assumes q1: "Quotient R1 Abs1 Rep1"
       
   300   and     q2: "Quotient R2 Abs2 Rep2"
       
   301   shows "f = (Rep1 ---> Abs2) ???"
       
   302 *)
       
   303 
       
   304 lemma LAMBDA_REP_ABS_RSP:
   542 lemma LAMBDA_REP_ABS_RSP:
   305   assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))"
   543   assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))"
   306   and     r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))"
   544   and     r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))"
   307   shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))"
   545   shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))"
   308 using r1 r2 by auto
   546 using r1 r2 by auto
   309 
   547 
   310 lemma REP_ABS_RSP:
   548 (* Not used *)
       
   549 lemma rep_abs_rsp_left:
   311   assumes q: "Quotient R Abs Rep"
   550   assumes q: "Quotient R Abs Rep"
   312   and     a: "R x1 x2"
   551   and     a: "R x1 x2"
   313   shows "R x1 (Rep (Abs x2))"
   552   shows "R x1 (Rep (Abs x2))"
   314 using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
   553 using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
   315 
   554 
   316 (* Not used *)
       
   317 lemma REP_ABS_RSP_LEFT:
       
   318   assumes q: "Quotient R Abs Rep"
       
   319   and     a: "R x1 x2"
       
   320   shows "R x1 (Rep (Abs x2))"
       
   321 using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q])
       
   322 
       
   323 (* ----------------------------------------------------- *)
       
   324 (* Quantifiers: FORALL, EXISTS, EXISTS_UNIQUE,           *)
       
   325 (*              Ball, Bex, RES_EXISTS_EQUIV              *)
       
   326 (* ----------------------------------------------------- *)
       
   327 
       
   328 (* bool theory: COND, LET *)
       
   329 
       
   330 lemma IF_PRS:
       
   331   assumes q: "Quotient R Abs Rep"
       
   332   shows "If a b c = Abs (If a (Rep b) (Rep c))"
       
   333 using Quotient_abs_rep[OF q] by auto
       
   334 
       
   335 (* ask peter: no use of q *)
       
   336 lemma IF_RSP:
       
   337   assumes q: "Quotient R Abs Rep"
       
   338   and     a: "a1 = a2" "R b1 b2" "R c1 c2"
       
   339   shows "R (If a1 b1 c1) (If a2 b2 c2)"
       
   340 using a by auto
       
   341 
       
   342 lemma LET_PRS:
       
   343   assumes q1: "Quotient R1 Abs1 Rep1"
       
   344   and     q2: "Quotient R2 Abs2 Rep2"
       
   345   shows "Let x f = Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f))"
       
   346 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
       
   347 
       
   348 lemma LET_RSP:
       
   349   assumes q1: "Quotient R1 Abs1 Rep1"
       
   350   and     q2: "Quotient R2 Abs2 Rep2"
       
   351   and     a1: "(R1 ===> R2) f g"
       
   352   and     a2: "R1 x y"
       
   353   shows "R2 (Let x f) (Let y g)"
       
   354 using fun_rel_MP[OF q1 q2 a1] a2
       
   355 by auto
       
   356 
       
   357 
       
   358 (* ask peter what are literal_case *)
       
   359 (* literal_case_PRS *)
       
   360 (* literal_case_RSP *)
       
   361 
       
   362 
       
   363 (* FUNCTION APPLICATION *)
       
   364 
       
   365 (* Not used *)
       
   366 lemma APPLY_PRS:
       
   367   assumes q1: "Quotient R1 Abs1 Rep1"
       
   368   and     q2: "Quotient R2 Abs2 Rep2"
       
   369   shows "f x = Abs2 (((Abs1 ---> Rep2) f) (Rep1 x))"
       
   370 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
       
   371 
       
   372 (* In the following theorem R1 can be instantiated with anything,
       
   373    but we know some of the types of the Rep and Abs functions;
       
   374    so by solving Quotient assumptions we can get a unique R2 that
       
   375    will be provable; which is why we need to use APPLY_RSP *)
       
   376 lemma apply_rsp:
       
   377   assumes q: "Quotient R1 Abs1 Rep1"
       
   378   and     a: "(R1 ===> R2) f g" "R1 x y"
       
   379   shows "R2 ((f::'a\<Rightarrow>'c) x) ((g::'a\<Rightarrow>'c) y)"
       
   380 using a by (rule fun_rel_IMP)
       
   381 
       
   382 lemma apply_rsp':
       
   383   assumes a: "(R1 ===> R2) f g" "R1 x y"
       
   384   shows "R2 (f x) (g y)"
       
   385 using a by (rule fun_rel_IMP)
       
   386 
       
   387 
       
   388 (* combinators: I, K, o, C, W *)
       
   389 
       
   390 (* We use id_simps which includes id_apply; so these 2 theorems can be removed *)
       
   391 lemma I_PRS:
       
   392   assumes q: "Quotient R Abs Rep"
       
   393   shows "id e = Abs (id (Rep e))"
       
   394 using Quotient_abs_rep[OF q] by auto
       
   395 
       
   396 lemma I_RSP:
       
   397   assumes q: "Quotient R Abs Rep"
       
   398   and     a: "R e1 e2"
       
   399   shows "R (id e1) (id e2)"
       
   400 using a by auto
       
   401 
       
   402 lemma o_PRS:
       
   403   assumes q1: "Quotient R1 Abs1 Rep1"
       
   404   and     q2: "Quotient R2 Abs2 Rep2"
       
   405   and     q3: "Quotient R3 Abs3 Rep3"
       
   406   shows "f o g = (Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g))"
       
   407 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
       
   408 unfolding o_def expand_fun_eq
       
   409 by simp
       
   410 
       
   411 lemma o_RSP:
       
   412   assumes q1: "Quotient R1 Abs1 Rep1"
       
   413   and     q2: "Quotient R2 Abs2 Rep2"
       
   414   and     q3: "Quotient R3 Abs3 Rep3"
       
   415   and     a1: "(R2 ===> R3) f1 f2"
       
   416   and     a2: "(R1 ===> R2) g1 g2"
       
   417   shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
       
   418 using a1 a2 unfolding o_def expand_fun_eq
       
   419 by (auto)
       
   420 
       
   421 
       
   422 
       
   423 
       
   424 
       
   425 lemma COND_PRS:
       
   426   assumes a: "Quotient R absf repf"
       
   427   shows "(if a then b else c) = absf (if a then repf b else repf c)"
       
   428   using a unfolding Quotient_def by auto
       
   429 
       
   430 
       
   431 
       
   432 
       
   433 
       
   434 (* Set of lemmas for regularisation of ball and bex *)
       
   435 lemma ball_reg_eqv:
       
   436   fixes P :: "'a \<Rightarrow> bool"
       
   437   assumes a: "equivp R"
       
   438   shows "Ball (Respects R) P = (All P)"
       
   439   by (metis equivp_def IN_RESPECTS a)
       
   440 
       
   441 lemma bex_reg_eqv:
       
   442   fixes P :: "'a \<Rightarrow> bool"
       
   443   assumes a: "equivp R"
       
   444   shows "Bex (Respects R) P = (Ex P)"
       
   445   by (metis equivp_def IN_RESPECTS a)
       
   446 
       
   447 lemma ball_reg_right:
       
   448   assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x"
       
   449   shows "All P \<longrightarrow> Ball R Q"
       
   450   by (metis COMBC_def Collect_def Collect_mem_eq a)
       
   451 
       
   452 lemma bex_reg_left:
       
   453   assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x"
       
   454   shows "Bex R Q \<longrightarrow> Ex P"
       
   455   by (metis COMBC_def Collect_def Collect_mem_eq a)
       
   456 
       
   457 lemma ball_reg_left:
       
   458   assumes a: "equivp R"
       
   459   shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P"
       
   460   by (metis equivp_reflp IN_RESPECTS a)
       
   461 
       
   462 lemma bex_reg_right:
       
   463   assumes a: "equivp R"
       
   464   shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P"
       
   465   by (metis equivp_reflp IN_RESPECTS a)
       
   466 
       
   467 lemma ball_reg_eqv_range:
       
   468   fixes P::"'a \<Rightarrow> bool"
       
   469   and x::"'a"
       
   470   assumes a: "equivp R2"
       
   471   shows   "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
       
   472   apply(rule iffI)
       
   473   apply(rule allI)
       
   474   apply(drule_tac x="\<lambda>y. f x" in bspec)
       
   475   apply(simp add: Respects_def IN_RESPECTS)
       
   476   apply(rule impI)
       
   477   using a equivp_reflp_symp_transp[of "R2"]
       
   478   apply(simp add: reflp_def)
       
   479   apply(simp)
       
   480   apply(simp)
       
   481   done
       
   482 
       
   483 lemma bex_reg_eqv_range:
       
   484   fixes P::"'a \<Rightarrow> bool"
       
   485   and x::"'a"
       
   486   assumes a: "equivp R2"
       
   487   shows   "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))"
       
   488   apply(auto)
       
   489   apply(rule_tac x="\<lambda>y. f x" in bexI)
       
   490   apply(simp)
       
   491   apply(simp add: Respects_def IN_RESPECTS)
       
   492   apply(rule impI)
       
   493   using a equivp_reflp_symp_transp[of "R2"]
       
   494   apply(simp add: reflp_def)
       
   495   done
       
   496 
       
   497 lemma all_reg:
       
   498   assumes a: "!x :: 'a. (P x --> Q x)"
       
   499   and     b: "All P"
       
   500   shows "All Q"
       
   501   using a b by (metis)
       
   502 
       
   503 lemma ex_reg:
       
   504   assumes a: "!x :: 'a. (P x --> Q x)"
       
   505   and     b: "Ex P"
       
   506   shows "Ex Q"
       
   507   using a b by (metis)
       
   508 
       
   509 lemma ball_reg:
       
   510   assumes a: "!x :: 'a. (R x --> P x --> Q x)"
       
   511   and     b: "Ball R P"
       
   512   shows "Ball R Q"
       
   513   using a b by (metis COMBC_def Collect_def Collect_mem_eq)
       
   514 
       
   515 lemma bex_reg:
       
   516   assumes a: "!x :: 'a. (R x --> P x --> Q x)"
       
   517   and     b: "Bex R P"
       
   518   shows "Bex R Q"
       
   519   using a b by (metis COMBC_def Collect_def Collect_mem_eq)
       
   520 
       
   521 lemma ball_all_comm:
       
   522   "(\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)) \<Longrightarrow> ((\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y))"
       
   523 by auto
       
   524 
       
   525 lemma bex_ex_comm:
       
   526   "((\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)) \<Longrightarrow> ((\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y))"
       
   527 by auto
       
   528 
       
   529 (* 2 lemmas needed for proving repabs_inj *)
       
   530 lemma ball_rsp:
       
   531   assumes a: "(R ===> (op =)) f g"
       
   532   shows "Ball (Respects R) f = Ball (Respects R) g"
       
   533   using a by (simp add: Ball_def IN_RESPECTS)
       
   534 
       
   535 lemma bex_rsp:
       
   536   assumes a: "(R ===> (op =)) f g"
       
   537   shows "(Bex (Respects R) f = Bex (Respects R) g)"
       
   538   using a by (simp add: Bex_def IN_RESPECTS)
       
   539 
       
   540 (* 2 lemmas needed for cleaning of quantifiers *)
       
   541 lemma all_prs:
       
   542   assumes a: "Quotient R absf repf"
       
   543   shows "Ball (Respects R) ((absf ---> id) f) = All f"
       
   544   using a unfolding Quotient_def
       
   545   by (metis IN_RESPECTS fun_map.simps id_apply)
       
   546 
       
   547 lemma ex_prs:
       
   548   assumes a: "Quotient R absf repf"
       
   549   shows "Bex (Respects R) ((absf ---> id) f) = Ex f"
       
   550   using a unfolding Quotient_def
       
   551   by (metis COMBC_def Collect_def Collect_mem_eq IN_RESPECTS fun_map.simps id_apply)
       
   552 
       
   553 
       
   554 (* UNUSED *)
       
   555 lemma Quotient_rel_abs:
       
   556   assumes a: "Quotient E Abs Rep"
       
   557   shows "E r s \<Longrightarrow> Abs r = Abs s"
       
   558 using a unfolding Quotient_def
       
   559 by blast
       
   560 
       
   561 lemma Quotient_rel_abs_eq:
       
   562   assumes a: "Quotient E Abs Rep"
       
   563   shows "E r r \<Longrightarrow> E s s \<Longrightarrow> E r s = (Abs r = Abs s)"
       
   564 using a unfolding Quotient_def
       
   565 by blast
       
   566 
       
   567 
       
   568 
   555 
   569 end
   556 end
   570 
   557