diff -r 143f58bb34f9 -r 0909deb8059b src/HOL/Quotient_Examples/Lift_Fun.thy --- a/src/HOL/Quotient_Examples/Lift_Fun.thy Thu May 26 16:57:14 2016 +0200 +++ b/src/HOL/Quotient_Examples/Lift_Fun.thy Thu May 26 17:51:22 2016 +0200 @@ -2,21 +2,21 @@ Author: Ondrej Kuncar *) -section {* Example of lifting definitions with contravariant or co/contravariant type variables *} +section \Example of lifting definitions with contravariant or co/contravariant type variables\ theory Lift_Fun imports Main "~~/src/HOL/Library/Quotient_Syntax" begin -text {* This file is meant as a test case. +text \This file is meant as a test case. 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. *} + type variables or type variables which are covariant and contravariant in the same time.\ -subsection {* Contravariant type variables *} +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. *} +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 =" @@ -36,10 +36,10 @@ quotient_definition "bij_betw' :: ('a \ 'b) \ 'a set \ 'b set \ bool" is bij_betw done -subsection {* Co/Contravariant type variables *} +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. *} +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) @@ -49,7 +49,7 @@ quotient_definition "map_endofun :: ('a \ 'b) \ ('b \ 'a) \ 'a endofun \ 'b endofun" is map_endofun' done -text {* Registration of the map function for 'a endofun. *} +text \Registration of the map function for 'a endofun.\ functor map_endofun : map_endofun proof - @@ -63,7 +63,7 @@ by (auto simp: map_endofun_def map_endofun'_def map_fun_def fun_eq_iff) (simp add: a o_assoc) qed -text {* Relator for 'a endofun. *} +text \Relator for 'a endofun.\ definition rel_endofun' :: "('a \ 'b \ bool) \ ('a \ 'a) \ ('b \ 'b) \ bool" @@ -110,8 +110,8 @@ 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. *} +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 done