# HG changeset patch # User huffman # Date 1243013662 25200 # Node ID 50deb3badfba112397c17ad8d796869d98fee55e # Parent 8a890890d14336ea75c6c3e3e16f5bd565c71501 add combinators for building copy functions diff -r 8a890890d143 -r 50deb3badfba src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Thu May 21 18:49:14 2009 -0700 +++ b/src/HOLCF/Domain.thy Fri May 22 10:34:22 2009 -0700 @@ -202,6 +202,65 @@ lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3 +subsection {* Combinators for building copy functions *} + +definition + cfun_fun :: "('b \ 'a) \ ('c \ 'd) \ ('a \ 'c) \ ('b \ 'd)" +where + "cfun_fun = (\ f g p. g oo p oo f)" + +definition + ssum_fun :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd" +where + "ssum_fun = (\ f g. sscase\(sinl oo f)\(sinr oo g))" + +definition + sprod_fun :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd" +where + "sprod_fun = (\ f g. ssplit\(\ x y. (:f\x, g\y:)))" + +definition + cprod_fun :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd" +where + "cprod_fun = (\ f g. csplit\(\ x y. x, g\y>))" + +definition + u_fun :: "('a \ 'b) \ 'a u \ 'b u" +where + "u_fun = (\ f. fup\(up oo f))" + +lemma cfun_fun_strict: "b\\ = \ \ cfun_fun\a\b\\ = \" +unfolding cfun_fun_def expand_cfun_eq by simp + +lemma ssum_fun_strict: "ssum_fun\a\b\\ = \" +unfolding ssum_fun_def by simp + +lemma sprod_fun_strict: "sprod_fun\a\b\\ = \" +unfolding sprod_fun_def by simp + +lemma u_fun_strict: "u_fun\a\\ = \" +unfolding u_fun_def by simp + +lemma ssum_fun_sinl: "x \ \ \ ssum_fun\f\g\(sinl\x) = sinl\(f\x)" +by (simp add: ssum_fun_def) + +lemma ssum_fun_sinr: "x \ \ \ ssum_fun\f\g\(sinr\x) = sinr\(g\x)" +by (simp add: ssum_fun_def) + +lemma sprod_fun_spair: + "x \ \ \ y \ \ \ sprod_fun\f\g\(:x, y:) = (:f\x, g\y:)" +by (simp add: sprod_fun_def) + +lemma u_fun_up: "u_fun\a\(up\x) = up\(a\x)" +by (simp add: u_fun_def) + +lemmas domain_fun_stricts = + ssum_fun_strict sprod_fun_strict u_fun_strict + +lemmas domain_fun_simps = + ssum_fun_sinl ssum_fun_sinr sprod_fun_spair u_fun_up + + subsection {* Installing the domain package *} lemmas con_strict_rules =