diff -r 8631dfe017a8 -r 28414668b43d src/HOLCF/Cprod.thy --- a/src/HOLCF/Cprod.thy Tue Oct 11 23:23:39 2005 +0200 +++ b/src/HOLCF/Cprod.thy Tue Oct 11 23:27:14 2005 +0200 @@ -182,31 +182,41 @@ subsection {* Continuous versions of constants *} -consts - cpair :: "'a \ 'b \ ('a * 'b)" (* continuous pairing *) - cfst :: "('a * 'b) \ 'a" - csnd :: "('a * 'b) \ 'b" +constdefs + cpair :: "'a \ 'b \ ('a * 'b)" (* continuous pairing *) + "cpair \ (\ x y. (x, y))" + + cfst :: "('a * 'b) \ 'a" + "cfst \ (\ p. fst p)" + + csnd :: "('a * 'b) \ 'b" + "csnd \ (\ p. snd p)" + csplit :: "('a \ 'b \ 'c) \ ('a * 'b) \ 'c" + "csplit \ (\ f p. f\(cfst\p)\(csnd\p))" syntax - "@ctuple" :: "['a, args] \ 'a * 'b" ("(1<_,/ _>)") + "_ctuple" :: "['a, args] \ 'a * 'b" ("(1<_,/ _>)") + +syntax (xsymbols) + "_ctuple" :: "['a, args] \ 'a * 'b" ("(1\_,/ _\)") translations "" == ">" "" == "cpair$x$y" +text {* syntax for @{text "LAM .e"} *} + syntax "_ctuple_pttrn" :: "[pttrn, patterns] => pttrn" ("<_,/ _>") +syntax (xsymbols) + "_ctuple_pttrn" :: "[pttrn, patterns] => pttrn" ("\_,/ _\") + translations "LAM . t" == "csplit$(LAM x . t)" "LAM . t" == "csplit$(LAM x y. t)" -defs - cpair_def: "cpair \ (\ x y. (x, y))" - cfst_def: "cfst \ (\ p. fst p)" - csnd_def: "csnd \ (\ p. snd p)" - csplit_def: "csplit \ (\ f p. f\(cfst\p)\(csnd\p))" subsection {* Convert all lemmas to the continuous versions *}