# HG changeset patch # User wenzelm # Date 1406195655 -7200 # Node ID dc59f147b27d2ed2c8a9ebef53dbb66017f82761 # Parent 0a28cf866d5d94efd8f83fdd7517f904f55767df more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy); diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/BNF_Def.thy --- a/src/HOL/BNF_Def.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/BNF_Def.thy Thu Jul 24 11:54:15 2014 +0200 @@ -53,21 +53,21 @@ lemma collect_comp: "collect F \ g = collect ((\f. f \ g) ` F)" by (rule ext) (auto simp only: comp_apply collect_def) -definition convol ("<_ , _>") where -" \ %a. (f a, g a)" +definition convol ("\(_,/ _)\") where +"\f, g\ \ \a. (f a, g a)" lemma fst_convol: -"fst \ = f" +"fst \ \f, g\ = f" apply(rule ext) unfolding convol_def by simp lemma snd_convol: -"snd \ = g" +"snd \ \f, g\ = g" apply(rule ext) unfolding convol_def by simp lemma convol_mem_GrpI: -"x \ A \ x \ (Collect (split (Grp A g)))" +"x \ A \ \id, g\ x \ (Collect (split (Grp A g)))" unfolding convol_def Grp_def by auto definition csquare where @@ -182,7 +182,7 @@ lemma rel_fun_iff_leq_vimage2p: "(rel_fun R S) f g = (R \ vimage2p f g S)" unfolding rel_fun_def vimage2p_def by auto -lemma convol_image_vimage2p: " fst, g \ snd> ` Collect (split (vimage2p f g R)) \ Collect (split R)" +lemma convol_image_vimage2p: "\f \ fst, g \ snd\ ` Collect (split (vimage2p f g R)) \ Collect (split R)" unfolding vimage2p_def convol_def by auto lemma vimage2p_Grp: "vimage2p f g P = Grp UNIV f OO P OO (Grp UNIV g)\\" diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy --- a/src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/BNF_Examples/Derivation_Trees/Prelim.thy Thu Jul 24 11:54:15 2014 +0200 @@ -11,11 +11,11 @@ imports "~~/src/HOL/Library/FSet" begin -notation BNF_Def.convol ("<_ , _>") +notation BNF_Def.convol ("\(_,/ _)\") declare fset_to_fset[simp] -lemma fst_snd_convol_o[simp]: " = s" +lemma fst_snd_convol_o[simp]: "\fst o s, snd o s\ = s" apply(rule ext) by (simp add: convol_def) abbreviation sm_abbrev (infix "\" 60) diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/BNF_Examples/Stream_Processor.thy --- a/src/HOL/BNF_Examples/Stream_Processor.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/BNF_Examples/Stream_Processor.thy Thu Jul 24 11:54:15 2014 +0200 @@ -152,10 +152,10 @@ bnf_axiomatization ('a, 'b) F for map: F -notation BNF_Def.convol ("<_ , _>") +notation BNF_Def.convol ("\(_,/ _)\") definition \ :: "('p,'a) F * 'b \ ('p,'a * 'b) F" where - "\ xb = F id a. (snd xb)> (fst xb)" + "\ xb = F id \id, \ a. (snd xb)\ (fst xb)" (* The strength laws for \: *) lemma \_natural: "F id (map_prod f g) o \ = \ o map_prod (F id f) g" diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/BNF_FP_Base.thy --- a/src/HOL/BNF_FP_Base.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/BNF_FP_Base.thy Thu Jul 24 11:54:15 2014 +0200 @@ -113,13 +113,13 @@ lemma rewriteL_comp_comp2: "\f \ g = l1 \ l2; l2 \ h = r\ \ f \ (g \ h) = l1 \ r" unfolding comp_def fun_eq_iff by auto -lemma convol_o: " \ h = h, g \ h>" +lemma convol_o: "\f, g\ \ h = \f \ h, g \ h\" unfolding convol_def by auto -lemma map_prod_o_convol: "map_prod h1 h2 \ =

f, h2 \ g>" +lemma map_prod_o_convol: "map_prod h1 h2 \ \f, g\ = \h1 \ f, h2 \ g\" unfolding convol_def by auto -lemma map_prod_o_convol_id: "(map_prod f id \ ) x = f , g> x" +lemma map_prod_o_convol_id: "(map_prod f id \ \id, g\) x = \id \ f, g\ x" unfolding map_prod_o_convol id_comp comp_id .. lemma o_case_sum: "h \ case_sum f g = case_sum (h \ f) (h \ g)" diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/BNF_LFP.thy --- a/src/HOL/BNF_LFP.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/BNF_LFP.thy Thu Jul 24 11:54:15 2014 +0200 @@ -41,23 +41,23 @@ lemma FieldI2: "(i, j) \ R \ j \ Field R" unfolding Field_def by auto -lemma fst_convol': "fst ( x) = f x" +lemma fst_convol': "fst (\f, g\ x) = f x" using fst_convol unfolding convol_def by simp -lemma snd_convol': "snd ( x) = g x" +lemma snd_convol': "snd (\f, g\ x) = g x" using snd_convol unfolding convol_def by simp -lemma convol_expand_snd: "fst o f = g \ = f" +lemma convol_expand_snd: "fst o f = g \ \g, snd o f\ = f" unfolding convol_def by auto lemma convol_expand_snd': assumes "(fst o f = g)" - shows "h = snd o f \ = f" + shows "h = snd o f \ \g, h\ = f" proof - - from assms have *: " = f" by (rule convol_expand_snd) - then have "h = snd o f \ h = snd o " by simp + from assms have *: "\g, snd o f\ = f" by (rule convol_expand_snd) + then have "h = snd o f \ h = snd o \g, snd o f\" by simp moreover have "\ \ h = snd o f" by (simp add: snd_convol) - moreover have "\ \ = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff) + moreover have "\ \ \g, h\ = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff) ultimately show ?thesis by simp qed lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" diff -r 0a28cf866d5d -r dc59f147b27d src/HOL/Main.thy --- a/src/HOL/Main.thy Thu Jul 24 11:51:22 2014 +0200 +++ b/src/HOL/Main.thy Thu Jul 24 11:54:15 2014 +0200 @@ -26,7 +26,7 @@ csum (infixr "+c" 65) and cprod (infixr "*c" 80) and cexp (infixr "^c" 90) and - convol ("<_ , _>") + convol ("\(_,/ _)\") hide_const (open) czero cinfinite cfinite csum cone ctwo Csum cprod cexp