more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.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 \<circ> g = collect ((\<lambda>f. f \<circ> g) ` F)"
by (rule ext) (auto simp only: comp_apply collect_def)
-definition convol ("<_ , _>") where
-"<f , g> \<equiv> %a. (f a, g a)"
+definition convol ("\<langle>(_,/ _)\<rangle>") where
+"\<langle>f, g\<rangle> \<equiv> \<lambda>a. (f a, g a)"
lemma fst_convol:
-"fst \<circ> <f , g> = f"
+"fst \<circ> \<langle>f, g\<rangle> = f"
apply(rule ext)
unfolding convol_def by simp
lemma snd_convol:
-"snd \<circ> <f , g> = g"
+"snd \<circ> \<langle>f, g\<rangle> = g"
apply(rule ext)
unfolding convol_def by simp
lemma convol_mem_GrpI:
-"x \<in> A \<Longrightarrow> <id , g> x \<in> (Collect (split (Grp A g)))"
+"x \<in> A \<Longrightarrow> \<langle>id, g\<rangle> x \<in> (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 \<le> vimage2p f g S)"
unfolding rel_fun_def vimage2p_def by auto
-lemma convol_image_vimage2p: "<f \<circ> fst, g \<circ> snd> ` Collect (split (vimage2p f g R)) \<subseteq> Collect (split R)"
+lemma convol_image_vimage2p: "\<langle>f \<circ> fst, g \<circ> snd\<rangle> ` Collect (split (vimage2p f g R)) \<subseteq> 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)\<inverse>\<inverse>"
--- 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 ("\<langle>(_,/ _)\<rangle>")
declare fset_to_fset[simp]
-lemma fst_snd_convol_o[simp]: "<fst o s, snd o s> = s"
+lemma fst_snd_convol_o[simp]: "\<langle>fst o s, snd o s\<rangle> = s"
apply(rule ext) by (simp add: convol_def)
abbreviation sm_abbrev (infix "\<oplus>" 60)
--- 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 ("\<langle>(_,/ _)\<rangle>")
definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
- "\<theta> xb = F id <id, \<lambda> a. (snd xb)> (fst xb)"
+ "\<theta> xb = F id \<langle>id, \<lambda> a. (snd xb)\<rangle> (fst xb)"
(* The strength laws for \<theta>: *)
lemma \<theta>_natural: "F id (map_prod f g) o \<theta> = \<theta> o map_prod (F id f) g"
--- 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: "\<lbrakk>f \<circ> g = l1 \<circ> l2; l2 \<circ> h = r\<rbrakk> \<Longrightarrow> f \<circ> (g \<circ> h) = l1 \<circ> r"
unfolding comp_def fun_eq_iff by auto
-lemma convol_o: "<f, g> \<circ> h = <f \<circ> h, g \<circ> h>"
+lemma convol_o: "\<langle>f, g\<rangle> \<circ> h = \<langle>f \<circ> h, g \<circ> h\<rangle>"
unfolding convol_def by auto
-lemma map_prod_o_convol: "map_prod h1 h2 \<circ> <f, g> = <h1 \<circ> f, h2 \<circ> g>"
+lemma map_prod_o_convol: "map_prod h1 h2 \<circ> \<langle>f, g\<rangle> = \<langle>h1 \<circ> f, h2 \<circ> g\<rangle>"
unfolding convol_def by auto
-lemma map_prod_o_convol_id: "(map_prod f id \<circ> <id , g>) x = <id \<circ> f , g> x"
+lemma map_prod_o_convol_id: "(map_prod f id \<circ> \<langle>id, g\<rangle>) x = \<langle>id \<circ> f, g\<rangle> x"
unfolding map_prod_o_convol id_comp comp_id ..
lemma o_case_sum: "h \<circ> case_sum f g = case_sum (h \<circ> f) (h \<circ> g)"
--- 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) \<in> R \<Longrightarrow> j \<in> Field R"
unfolding Field_def by auto
-lemma fst_convol': "fst (<f, g> x) = f x"
+lemma fst_convol': "fst (\<langle>f, g\<rangle> x) = f x"
using fst_convol unfolding convol_def by simp
-lemma snd_convol': "snd (<f, g> x) = g x"
+lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x"
using snd_convol unfolding convol_def by simp
-lemma convol_expand_snd: "fst o f = g \<Longrightarrow> <g, snd o f> = f"
+lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f"
unfolding convol_def by auto
lemma convol_expand_snd':
assumes "(fst o f = g)"
- shows "h = snd o f \<longleftrightarrow> <g, h> = f"
+ shows "h = snd o f \<longleftrightarrow> \<langle>g, h\<rangle> = f"
proof -
- from assms have *: "<g, snd o f> = f" by (rule convol_expand_snd)
- then have "h = snd o f \<longleftrightarrow> h = snd o <g, snd o f>" by simp
+ from assms have *: "\<langle>g, snd o f\<rangle> = f" by (rule convol_expand_snd)
+ then have "h = snd o f \<longleftrightarrow> h = snd o \<langle>g, snd o f\<rangle>" by simp
moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol)
- moreover have "\<dots> \<longleftrightarrow> <g, h> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
+ moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = 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 \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
--- 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 ("\<langle>(_,/ _)\<rangle>")
hide_const (open)
czero cinfinite cfinite csum cone ctwo Csum cprod cexp