src/HOL/BNF_FP_Base.thy
changeset 57982 0a35354137a5
parent 57308 972f0aa7091b
child 57992 1f9ab71d43a5
     1.1 --- a/src/HOL/BNF_FP_Base.thy	Wed Apr 23 10:23:26 2014 +0200
     1.2 +++ b/src/HOL/BNF_FP_Base.thy	Wed Apr 23 10:23:26 2014 +0200
     1.3 @@ -89,34 +89,34 @@
     1.4  lemma spec2: "\<forall>x y. P x y \<Longrightarrow> P x y"
     1.5  by blast
     1.6  
     1.7 -lemma rewriteR_comp_comp: "\<lbrakk>g o h = r\<rbrakk> \<Longrightarrow> f o g o h = f o r"
     1.8 +lemma rewriteR_comp_comp: "\<lbrakk>g \<circ> h = r\<rbrakk> \<Longrightarrow> f \<circ> g \<circ> h = f \<circ> r"
     1.9    unfolding comp_def fun_eq_iff by auto
    1.10  
    1.11 -lemma rewriteR_comp_comp2: "\<lbrakk>g o h = r1 o r2; f o r1 = l\<rbrakk> \<Longrightarrow> f o g o h = l o r2"
    1.12 +lemma rewriteR_comp_comp2: "\<lbrakk>g \<circ> h = r1 \<circ> r2; f \<circ> r1 = l\<rbrakk> \<Longrightarrow> f \<circ> g \<circ> h = l \<circ> r2"
    1.13    unfolding comp_def fun_eq_iff by auto
    1.14  
    1.15 -lemma rewriteL_comp_comp: "\<lbrakk>f o g = l\<rbrakk> \<Longrightarrow> f o (g o h) = l o h"
    1.16 +lemma rewriteL_comp_comp: "\<lbrakk>f \<circ> g = l\<rbrakk> \<Longrightarrow> f \<circ> (g \<circ> h) = l \<circ> h"
    1.17    unfolding comp_def fun_eq_iff by auto
    1.18  
    1.19 -lemma rewriteL_comp_comp2: "\<lbrakk>f o g = l1 o l2; l2 o h = r\<rbrakk> \<Longrightarrow> f o (g o h) = l1 o r"
    1.20 +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"
    1.21    unfolding comp_def fun_eq_iff by auto
    1.22  
    1.23 -lemma convol_o: "<f, g> o h = <f o h, g o h>"
    1.24 +lemma convol_o: "<f, g> \<circ> h = <f \<circ> h, g \<circ> h>"
    1.25    unfolding convol_def by auto
    1.26  
    1.27 -lemma map_prod_o_convol: "map_prod h1 h2 o <f, g> = <h1 o f, h2 o g>"
    1.28 +lemma map_prod_o_convol: "map_prod h1 h2 \<circ> <f, g> = <h1 \<circ> f, h2 \<circ> g>"
    1.29    unfolding convol_def by auto
    1.30  
    1.31  lemma map_prod_o_convol_id: "(map_prod f id \<circ> <id , g>) x = <id \<circ> f , g> x"
    1.32    unfolding map_prod_o_convol id_comp comp_id ..
    1.33  
    1.34 -lemma o_case_sum: "h o case_sum f g = case_sum (h o f) (h o g)"
    1.35 +lemma o_case_sum: "h \<circ> case_sum f g = case_sum (h \<circ> f) (h \<circ> g)"
    1.36    unfolding comp_def by (auto split: sum.splits)
    1.37  
    1.38 -lemma case_sum_o_map_sum: "case_sum f g o map_sum h1 h2 = case_sum (f o h1) (g o h2)"
    1.39 +lemma case_sum_o_map_sum: "case_sum f g \<circ> map_sum h1 h2 = case_sum (f \<circ> h1) (g \<circ> h2)"
    1.40    unfolding comp_def by (auto split: sum.splits)
    1.41  
    1.42 -lemma case_sum_o_map_sum_id: "(case_sum id g o map_sum f id) x = case_sum (f o id) g x"
    1.43 +lemma case_sum_o_map_sum_id: "(case_sum id g \<circ> map_sum f id) x = case_sum (f \<circ> id) g x"
    1.44    unfolding case_sum_o_map_sum id_comp comp_id ..
    1.45  
    1.46  lemma rel_fun_def_butlast:
    1.47 @@ -144,7 +144,7 @@
    1.48  
    1.49  lemma
    1.50    assumes "type_definition Rep Abs UNIV"
    1.51 -  shows type_copy_Rep_o_Abs: "Rep \<circ> Abs = id" and type_copy_Abs_o_Rep: "Abs o Rep = id"
    1.52 +  shows type_copy_Rep_o_Abs: "Rep \<circ> Abs = id" and type_copy_Abs_o_Rep: "Abs \<circ> Rep = id"
    1.53    unfolding fun_eq_iff comp_apply id_apply
    1.54      type_definition.Abs_inverse[OF assms UNIV_I] type_definition.Rep_inverse[OF assms] by simp_all
    1.55  
    1.56 @@ -152,7 +152,7 @@
    1.57    assumes "type_definition Rep Abs UNIV"
    1.58            "type_definition Rep' Abs' UNIV"
    1.59            "type_definition Rep'' Abs'' UNIV"
    1.60 -  shows "Abs' o M o Rep'' = (Abs' o M1 o Rep) o (Abs o M2 o Rep'') \<Longrightarrow> M1 o M2 = M"
    1.61 +  shows "Abs' \<circ> M \<circ> Rep'' = (Abs' \<circ> M1 \<circ> Rep) \<circ> (Abs \<circ> M2 \<circ> Rep'') \<Longrightarrow> M1 \<circ> M2 = M"
    1.62    by (rule sym) (auto simp: fun_eq_iff type_definition.Abs_inject[OF assms(2) UNIV_I UNIV_I]
    1.63      type_definition.Abs_inverse[OF assms(1) UNIV_I]
    1.64      type_definition.Abs_inverse[OF assms(3) UNIV_I] dest: spec[of _ "Abs'' x" for x])
    1.65 @@ -160,7 +160,7 @@
    1.66  lemma vimage2p_id: "vimage2p id id R = R"
    1.67    unfolding vimage2p_def by auto
    1.68  
    1.69 -lemma vimage2p_comp: "vimage2p (f1 o f2) (g1 o g2) = vimage2p f2 g2 o vimage2p f1 g1"
    1.70 +lemma vimage2p_comp: "vimage2p (f1 \<circ> f2) (g1 \<circ> g2) = vimage2p f2 g2 \<circ> vimage2p f1 g1"
    1.71    unfolding fun_eq_iff vimage2p_def o_apply by simp
    1.72  
    1.73  ML_file "Tools/BNF/bnf_fp_util.ML"