From 2116cee735ca505e2eae260f48341a4d8ab24117 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Thu, 28 Feb 2019 14:15:31 +0100 Subject: more convenient syntax --- Prod.thy | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) (limited to 'Prod.thy') diff --git a/Prod.thy b/Prod.thy index d27c51e..48f0151 100644 --- a/Prod.thy +++ b/Prod.thy @@ -75,7 +75,12 @@ section \Function composition\ definition compose :: "[t, t, t] \ t" where "compose A g f \ \x: A. g`(f`x)" -syntax "_compose" :: "[t, t] \ t" (infixr "o" 110) +syntax "_compose" :: "[t, t, t] \ t" ("(2_ o[_]/ _)" [111, 0, 110] 110) +translations "g o[A] f" \ "(CONST compose) A g f" + +text \The composition @{term "g o[A] f"} is annotated with the domain @{term A} of @{term f}.\ + +syntax "_compose'" :: "[t, t] \ t" (infixr "o" 110) text \Pretty-printing switch for composition; hides domain type information.\ @@ -84,7 +89,7 @@ ML \val pretty_compose = Attrib.setup_config_bool @{binding "pretty_compos print_translation \ let fun compose_tr' ctxt [A, g, f] = if Config.get ctxt pretty_compose - then Syntax.const @{syntax_const "_compose"} $ g $ f + then Syntax.const @{syntax_const "_compose'"} $ g $ f else @{const compose} $ A $ g $ f in [(@{const_syntax compose}, compose_tr')] @@ -93,12 +98,12 @@ end lemma compose_assoc: assumes "A: U i" and "f: A \ B" and "g: B \ C" and "h: \x: C. D x" - shows "compose A (compose B h g) f \ compose A h (compose A g f)" + shows "(h o[B] g) o[A] f \ h o[A] g o[A] f" unfolding compose_def by (derive lems: assms cong) lemma compose_comp: assumes "A: U i" and "\x. x: A \ b x: B" and "\x. x: B \ c x: C x" - shows "compose A (\x: B. c x) (\x: A. b x) \ \x: A. c (b x)" + shows "(\x: B. c x) o[A] (\x: A. b x) \ \x: A. c (b x)" unfolding compose_def by (derive lems: assms cong) declare compose_comp [comp] -- cgit v1.2.3