# HG changeset patch # User kuncar # Date 1323436925 -3600 # Node ID 7fa9410c746a5ead7db72f1352c3ed00ef5d1f67 # Parent 2373d86cf2e85fd4afaa7e4891f9b44ac8390d6e added an example file with lifting of constants with contravariant and co/contravariant types diff -r 2373d86cf2e8 -r 7fa9410c746a src/HOL/Quotient_Examples/Lift_Fun.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Quotient_Examples/Lift_Fun.thy Fri Dec 09 14:22:05 2011 +0100 @@ -0,0 +1,82 @@ +(* Title: HOL/Quotient_Examples/Lift_Fun.thy + Author: Ondrej Kuncar +*) + +header {* Example of lifting definitions with contravariant or co/contravariant type variables *} + + +theory Lift_Fun +imports Main +begin + +text {* This file is meant as a test case for features introduced in the changeset 2d8949268303. + It contains examples of lifting definitions with quotients that have contravariant + type variables or type variables which are covariant and contravariant in the same time. *} + +subsection {* Contravariant type variables *} + +text {* 'a is a contravariant type variable and we are able to map over this variable + in the following four definitions. This example is based on HOL/Fun.thy. *} + +quotient_type +('a, 'b) fun' (infixr "\" 55) = "'a \ 'b" / "op =" + by (simp add: identity_equivp) + +quotient_definition "comp' :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" is + "comp :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" + +quotient_definition "fcomp' :: ('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" is + fcomp + +quotient_definition "map_fun' :: ('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" + is "map_fun::('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" + +quotient_definition "inj_on' :: ('a \ 'b) \ 'a set \ bool" is inj_on + +quotient_definition "bij_betw' :: ('a \ 'b) \ 'a set \ 'b set \ bool" is bij_betw + + +subsection {* Co/Contravariant type variables *} + +text {* 'a is a covariant and contravariant type variable in the same time. + The following example is a bit artificial. We haven't had a natural one yet. *} + +quotient_type 'a endofun = "'a \ 'a" / "op =" by (simp add: identity_equivp) + +definition map_endofun' :: "('a \ 'b) \ ('b \ 'a) \ ('a => 'a) \ ('b => 'b)" + where "map_endofun' f g e = map_fun g f e" + +quotient_definition "map_endofun :: ('a \ 'b) \ ('b \ 'a) \ 'a endofun \ 'b endofun" is + map_endofun' + +text {* Registration of the map function for 'a endofun. *} + +enriched_type map_endofun : map_endofun +proof - + have "\ x. abs_endofun (rep_endofun x) = x" using Quotient_endofun by (auto simp: Quotient_def) + then show "map_endofun id id = id" + by (auto simp: map_endofun_def map_endofun'_def map_fun_def fun_eq_iff) + + have a:"\ x. rep_endofun (abs_endofun x) = x" using Quotient_endofun + Quotient_rep_abs[of "(op =)" abs_endofun rep_endofun] by blast + show "\f g h i. map_endofun f g \ map_endofun h i = map_endofun (f \ h) (i \ g)" + by (auto simp: map_endofun_def map_endofun'_def map_fun_def fun_eq_iff) (simp add: a o_assoc) +qed + +quotient_definition "endofun_id_id :: ('a endofun) endofun" is "id :: ('a \ 'a) \ ('a \ 'a)" + +term endofun_id_id +thm endofun_id_id_def + +quotient_type 'a endofun' = "'a endofun" / "op =" by (simp add: identity_equivp) + +text {* We have to map "'a endofun" to "('a endofun') endofun", i.e., mapping (lifting) + over a type variable which is a covariant and contravariant type variable. *} + +quotient_definition "endofun'_id_id :: ('a endofun') endofun'" is endofun_id_id + +term endofun'_id_id +thm endofun'_id_id_def + + +end