Renamed (^) to [^] in preparation of the move from "op X" to (X)
authornipkow
Fri, 05 Jan 2018 18:41:42 +0100
changeset 67341 df79ef3b3a41
parent 67340 150d40a25622
child 67342 7905adb28bdc
Renamed (^) to [^] in preparation of the move from "op X" to (X)
NEWS
src/HOL/Algebra/FiniteProduct.thy
src/HOL/Algebra/Group.thy
src/HOL/Algebra/More_Finite_Product.thy
src/HOL/Algebra/More_Group.thy
src/HOL/Algebra/Multiplicative_Group.thy
src/HOL/Algebra/Ring.thy
src/HOL/Algebra/UnivPoly.thy
src/HOL/Decision_Procs/Algebra_Aux.thy
src/HOL/Decision_Procs/Commutative_Ring.thy
src/HOL/Decision_Procs/Reflective_Field.thy
src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy
src/HOL/Number_Theory/Prime_Powers.thy
src/HOL/Number_Theory/Residues.thy
src/HOL/Quotient_Examples/Int_Pow.thy
--- a/NEWS	Fri Jan 05 15:24:57 2018 +0100
+++ b/NEWS	Fri Jan 05 18:41:42 2018 +0100
@@ -166,6 +166,8 @@
 * Predicate pairwise_coprime abolished, use "pairwise coprime" instead.
 INCOMPATIBILITY.
 
+* HOL-Algebra: renamed (^) to [^]
+
 * Session HOL-Analysis: Moebius functions and the Riemann mapping
 theorem.
 
--- a/src/HOL/Algebra/FiniteProduct.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/FiniteProduct.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -502,12 +502,12 @@
 
 lemma finprod_const:
   assumes a [simp]: "a : carrier G"
-    shows "finprod G (%x. a) A = a (^) card A"
+    shows "finprod G (%x. a) A = a [^] card A"
 proof (induct A rule: infinite_finite_induct)
   case (insert b A)
   show ?case 
   proof (subst finprod_insert[OF insert(1-2)])
-    show "a \<otimes> (\<Otimes>x\<in>A. a) = a (^) card (insert b A)"
+    show "a \<otimes> (\<Otimes>x\<in>A. a) = a [^] card (insert b A)"
       by (insert insert, auto, subst m_comm, auto)
   qed auto
 qed auto
--- a/src/HOL/Algebra/Group.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/Group.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -30,7 +30,7 @@
   where "Units G = {y. y \<in> carrier G \<and> (\<exists>x \<in> carrier G. x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> \<and> y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)}"
 
 consts
-  pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a"  (infixr "'(^')\<index>" 75)
+  pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a"  (infixr "[^]\<index>" 75)
 
 overloading nat_pow == "pow :: [_, 'a, nat] => 'a"
 begin
@@ -44,7 +44,7 @@
     in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))"
 end
 
-lemma int_pow_int: "x (^)\<^bsub>G\<^esub> (int n) = x (^)\<^bsub>G\<^esub> n"
+lemma int_pow_int: "x [^]\<^bsub>G\<^esub> (int n) = x [^]\<^bsub>G\<^esub> n"
 by(simp add: int_pow_def nat_pow_def)
 
 locale monoid =
@@ -196,27 +196,27 @@
 text \<open>Power\<close>
 
 lemma (in monoid) nat_pow_closed [intro, simp]:
-  "x \<in> carrier G ==> x (^) (n::nat) \<in> carrier G"
+  "x \<in> carrier G ==> x [^] (n::nat) \<in> carrier G"
   by (induct n) (simp_all add: nat_pow_def)
 
 lemma (in monoid) nat_pow_0 [simp]:
-  "x (^) (0::nat) = \<one>"
+  "x [^] (0::nat) = \<one>"
   by (simp add: nat_pow_def)
 
 lemma (in monoid) nat_pow_Suc [simp]:
-  "x (^) (Suc n) = x (^) n \<otimes> x"
+  "x [^] (Suc n) = x [^] n \<otimes> x"
   by (simp add: nat_pow_def)
 
 lemma (in monoid) nat_pow_one [simp]:
-  "\<one> (^) (n::nat) = \<one>"
+  "\<one> [^] (n::nat) = \<one>"
   by (induct n) simp_all
 
 lemma (in monoid) nat_pow_mult:
-  "x \<in> carrier G ==> x (^) (n::nat) \<otimes> x (^) m = x (^) (n + m)"
+  "x \<in> carrier G ==> x [^] (n::nat) \<otimes> x [^] m = x [^] (n + m)"
   by (induct m) (simp_all add: m_assoc [THEN sym])
 
 lemma (in monoid) nat_pow_pow:
-  "x \<in> carrier G ==> (x (^) n) (^) m = x (^) (n * m::nat)"
+  "x \<in> carrier G ==> (x [^] n) [^] m = x [^] (n * m::nat)"
   by (induct m) (simp, simp add: nat_pow_mult add.commute)
 
 
@@ -406,33 +406,33 @@
 text \<open>Power\<close>
 
 lemma (in group) int_pow_def2:
-  "a (^) (z::int) = (if z < 0 then inv (a (^) (nat (-z))) else a (^) (nat z))"
+  "a [^] (z::int) = (if z < 0 then inv (a [^] (nat (-z))) else a [^] (nat z))"
   by (simp add: int_pow_def nat_pow_def Let_def)
 
 lemma (in group) int_pow_0 [simp]:
-  "x (^) (0::int) = \<one>"
+  "x [^] (0::int) = \<one>"
   by (simp add: int_pow_def2)
 
 lemma (in group) int_pow_one [simp]:
-  "\<one> (^) (z::int) = \<one>"
+  "\<one> [^] (z::int) = \<one>"
   by (simp add: int_pow_def2)
 
 (* The following are contributed by Joachim Breitner *)
 
 lemma (in group) int_pow_closed [intro, simp]:
-  "x \<in> carrier G ==> x (^) (i::int) \<in> carrier G"
+  "x \<in> carrier G ==> x [^] (i::int) \<in> carrier G"
   by (simp add: int_pow_def2)
 
 lemma (in group) int_pow_1 [simp]:
-  "x \<in> carrier G \<Longrightarrow> x (^) (1::int) = x"
+  "x \<in> carrier G \<Longrightarrow> x [^] (1::int) = x"
   by (simp add: int_pow_def2)
 
 lemma (in group) int_pow_neg:
-  "x \<in> carrier G \<Longrightarrow> x (^) (-i::int) = inv (x (^) i)"
+  "x \<in> carrier G \<Longrightarrow> x [^] (-i::int) = inv (x [^] i)"
   by (simp add: int_pow_def2)
 
 lemma (in group) int_pow_mult:
-  "x \<in> carrier G \<Longrightarrow> x (^) (i + j::int) = x (^) i \<otimes> x (^) j"
+  "x \<in> carrier G \<Longrightarrow> x [^] (i + j::int) = x [^] i \<otimes> x [^] j"
 proof -
   have [simp]: "-i - j = -j - i" by simp
   assume "x : carrier G" then
@@ -441,7 +441,7 @@
 qed
 
 lemma (in group) int_pow_diff:
-  "x \<in> carrier G \<Longrightarrow> x (^) (n - m :: int) = x (^) n \<otimes> inv (x (^) m)"
+  "x \<in> carrier G \<Longrightarrow> x [^] (n - m :: int) = x [^] n \<otimes> inv (x [^] m)"
 by(simp only: diff_conv_add_uminus int_pow_mult int_pow_neg)
 
 lemma (in group) inj_on_multc: "c \<in> carrier G \<Longrightarrow> inj_on (\<lambda>x. x \<otimes> c) (carrier G)"
@@ -676,7 +676,7 @@
 
 (* Contributed by Joachim Breitner *)
 lemma (in group) int_pow_is_hom:
-  "x \<in> carrier G \<Longrightarrow> (op(^) x) \<in> hom \<lparr> carrier = UNIV, mult = op +, one = 0::int \<rparr> G "
+  "x \<in> carrier G \<Longrightarrow> (op[^] x) \<in> hom \<lparr> carrier = UNIV, mult = op +, one = 0::int \<rparr> G "
   unfolding hom_def by (simp add: int_pow_mult)
 
 
@@ -737,7 +737,7 @@
 
 lemma (in comm_monoid) nat_pow_distr:
   "[| x \<in> carrier G; y \<in> carrier G |] ==>
-  (x \<otimes> y) (^) (n::nat) = x (^) n \<otimes> y (^) n"
+  (x \<otimes> y) [^] (n::nat) = x [^] n \<otimes> y [^] n"
   by (induct n) (simp, simp add: m_ac)
 
 locale comm_group = comm_monoid + group
--- a/src/HOL/Algebra/More_Finite_Product.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/More_Finite_Product.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -82,14 +82,14 @@
 
 lemma (in monoid) units_of_pow:
   fixes n :: nat
-  shows "x \<in> Units G \<Longrightarrow> x (^)\<^bsub>units_of G\<^esub> n = x (^)\<^bsub>G\<^esub> n"
+  shows "x \<in> Units G \<Longrightarrow> x [^]\<^bsub>units_of G\<^esub> n = x [^]\<^bsub>G\<^esub> n"
   apply (induct n)
   apply (auto simp add: units_group group.is_monoid
     monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult)
   done
 
 lemma (in cring) units_power_order_eq_one:
-  "finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a (^) card(Units R) = \<one>"
+  "finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a [^] card(Units R) = \<one>"
   apply (subst units_of_carrier [symmetric])
   apply (subst units_of_one [symmetric])
   apply (subst units_of_pow [symmetric])
--- a/src/HOL/Algebra/More_Group.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/More_Group.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -113,14 +113,14 @@
 lemma (in comm_group) power_order_eq_one:
   assumes fin [simp]: "finite (carrier G)"
     and a [simp]: "a \<in> carrier G"
-  shows "a (^) card(carrier G) = one G"
+  shows "a [^] card(carrier G) = one G"
 proof -
   have "(\<Otimes>x\<in>carrier G. x) = (\<Otimes>x\<in>carrier G. a \<otimes> x)"
     by (subst (2) finprod_reindex [symmetric],
       auto simp add: Pi_def inj_on_const_mult surj_const_mult)
   also have "\<dots> = (\<Otimes>x\<in>carrier G. a) \<otimes> (\<Otimes>x\<in>carrier G. x)"
     by (auto simp add: finprod_multf Pi_def)
-  also have "(\<Otimes>x\<in>carrier G. a) = a (^) card(carrier G)"
+  also have "(\<Otimes>x\<in>carrier G. a) = a [^] card(carrier G)"
     by (auto simp add: finprod_const)
   finally show ?thesis
 (* uses the preceeding lemma *)
--- a/src/HOL/Algebra/Multiplicative_Group.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/Multiplicative_Group.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -84,7 +84,7 @@
 
 lemma evalRR_monom:
   assumes a: "a \<in> carrier R" and x: "x \<in> carrier R"
-  shows "eval R R id x (monom P a d) = a \<otimes> x (^) d"
+  shows "eval R R id x (monom P a d) = a \<otimes> x [^] d"
 proof -
   interpret UP_pre_univ_prop R R id by unfold_locales simp
   show ?thesis using assms by (simp add: eval_monom)
@@ -271,33 +271,33 @@
 lemma pow_eq_div2 :
   fixes m n :: nat
   assumes x_car: "x \<in> carrier G"
-  assumes pow_eq: "x (^) m = x (^) n"
-  shows "x (^) (m - n) = \<one>"
+  assumes pow_eq: "x [^] m = x [^] n"
+  shows "x [^] (m - n) = \<one>"
 proof (cases "m < n")
   case False
-  have "\<one> \<otimes> x (^) m = x (^) m" by (simp add: x_car)
-  also have "\<dots> = x (^) (m - n) \<otimes> x (^) n"
+  have "\<one> \<otimes> x [^] m = x [^] m" by (simp add: x_car)
+  also have "\<dots> = x [^] (m - n) \<otimes> x [^] n"
     using False by (simp add: nat_pow_mult x_car)
-  also have "\<dots> = x (^) (m - n) \<otimes> x (^) m"
+  also have "\<dots> = x [^] (m - n) \<otimes> x [^] m"
     by (simp add: pow_eq)
   finally show ?thesis by (simp add: x_car)
 qed simp
 
-definition ord where "ord a = Min {d \<in> {1 .. order G} . a (^) d = \<one>}"
+definition ord where "ord a = Min {d \<in> {1 .. order G} . a [^] d = \<one>}"
 
 lemma
   assumes finite:"finite (carrier G)"
   assumes a:"a \<in> carrier G"
   shows ord_ge_1: "1 \<le> ord a" and ord_le_group_order: "ord a \<le> order G"
-    and pow_ord_eq_1: "a (^) ord a = \<one>"
+    and pow_ord_eq_1: "a [^] ord a = \<one>"
 proof -
-  have "\<not>inj_on (\<lambda>x. a (^) x) {0 .. order G}"
+  have "\<not>inj_on (\<lambda>x. a [^] x) {0 .. order G}"
   proof (rule notI)
-    assume A: "inj_on (\<lambda>x. a (^) x) {0 .. order G}"
+    assume A: "inj_on (\<lambda>x. a [^] x) {0 .. order G}"
     have "order G + 1 = card {0 .. order G}" by simp
-    also have "\<dots> = card ((\<lambda>x. a (^) x) ` {0 .. order G})" (is "_ = card ?S")
+    also have "\<dots> = card ((\<lambda>x. a [^] x) ` {0 .. order G})" (is "_ = card ?S")
       using A by (simp add: card_image)
-    also have "?S = {a (^) x | x. x \<in> {0 .. order G}}" by blast
+    also have "?S = {a [^] x | x. x \<in> {0 .. order G}}" by blast
     also have "\<dots> \<subseteq> carrier G" (is "?S \<subseteq> _") using a by blast
     then have "card ?S \<le> order G" unfolding order_def
       by (rule card_mono[OF finite])
@@ -305,8 +305,8 @@
   qed
 
   then obtain x y where x_y:"x \<noteq> y" "x \<in> {0 .. order G}" "y \<in> {0 .. order G}"
-                        "a (^) x = a (^) y" unfolding inj_on_def by blast
-  obtain d where "1 \<le> d" "a (^) d = \<one>" "d \<le> order G"
+                        "a [^] x = a [^] y" unfolding inj_on_def by blast
+  obtain d where "1 \<le> d" "a [^] d = \<one>" "d \<le> order G"
   proof cases
     assume "y < x" with x_y show ?thesis
       by (intro that[where d="x - y"]) (auto simp add: pow_eq_div2[OF a])
@@ -314,22 +314,22 @@
     assume "\<not>y < x" with x_y show ?thesis
       by (intro that[where d="y - x"]) (auto simp add: pow_eq_div2[OF a])
   qed
-  hence "ord a \<in> {d \<in> {1 .. order G} . a (^) d = \<one>}"
-    unfolding ord_def using Min_in[of "{d \<in> {1 .. order G} . a (^) d = \<one>}"]
+  hence "ord a \<in> {d \<in> {1 .. order G} . a [^] d = \<one>}"
+    unfolding ord_def using Min_in[of "{d \<in> {1 .. order G} . a [^] d = \<one>}"]
     by fastforce
-  then show "1 \<le> ord a" and "ord a \<le> order G" and "a (^) ord a = \<one>"
+  then show "1 \<le> ord a" and "ord a \<le> order G" and "a [^] ord a = \<one>"
     by (auto simp: order_def)
 qed
 
 lemma finite_group_elem_finite_ord :
   assumes "finite (carrier G)" "x \<in> carrier G"
-  shows "\<exists> d::nat. d \<ge> 1 \<and> x (^) d = \<one>"
+  shows "\<exists> d::nat. d \<ge> 1 \<and> x [^] d = \<one>"
   using assms ord_ge_1 pow_ord_eq_1 by auto
 
 lemma ord_min:
-  assumes  "finite (carrier G)" "1 \<le> d" "a \<in> carrier G" "a (^) d = \<one>" shows "ord a \<le> d"
+  assumes  "finite (carrier G)" "1 \<le> d" "a \<in> carrier G" "a [^] d = \<one>" shows "ord a \<le> d"
 proof -
-  define Ord where "Ord = {d \<in> {1..order G}. a (^) d = \<one>}"
+  define Ord where "Ord = {d \<in> {1..order G}. a [^] d = \<one>}"
   have fin: "finite Ord" by (auto simp: Ord_def)
   have in_ord: "ord a \<in> Ord"
     using assms pow_ord_eq_1 ord_ge_1 ord_le_group_order by (auto simp: Ord_def)
@@ -350,22 +350,22 @@
 lemma ord_inj :
   assumes finite: "finite (carrier G)"
   assumes a: "a \<in> carrier G"
-  shows "inj_on (\<lambda> x . a (^) x) {0 .. ord a - 1}"
+  shows "inj_on (\<lambda> x . a [^] x) {0 .. ord a - 1}"
 proof (rule inj_onI, rule ccontr)
-  fix x y assume A: "x \<in> {0 .. ord a - 1}" "y \<in> {0 .. ord a - 1}" "a (^) x= a (^) y" "x \<noteq> y"
+  fix x y assume A: "x \<in> {0 .. ord a - 1}" "y \<in> {0 .. ord a - 1}" "a [^] x= a [^] y" "x \<noteq> y"
 
-  have "finite {d \<in> {1..order G}. a (^) d = \<one>}" by auto
+  have "finite {d \<in> {1..order G}. a [^] d = \<one>}" by auto
 
   { fix x y assume A: "x < y" "x \<in> {0 .. ord a - 1}" "y \<in> {0 .. ord a - 1}"
-        "a (^) x = a (^) y"
+        "a [^] x = a [^] y"
     hence "y - x < ord a" by auto
     also have "\<dots> \<le> order G" using assms by (simp add: ord_le_group_order)
     finally have y_x_range:"y - x \<in> {1 .. order G}" using A by force
-    have "a (^) (y-x) = \<one>" using a A by (simp add: pow_eq_div2)
+    have "a [^] (y-x) = \<one>" using a A by (simp add: pow_eq_div2)
 
-    hence y_x:"y - x \<in> {d \<in> {1.. order G}. a (^) d = \<one>}" using y_x_range by blast
+    hence y_x:"y - x \<in> {d \<in> {1.. order G}. a [^] d = \<one>}" using y_x_range by blast
     have "min (y - x) (ord a) = ord a"
-      using Min.in_idem[OF \<open>finite {d \<in> {1 .. order G} . a (^) d = \<one>}\<close> y_x] ord_def by auto
+      using Min.in_idem[OF \<open>finite {d \<in> {1 .. order G} . a [^] d = \<one>}\<close> y_x] ord_def by auto
     with \<open>y - x < ord a\<close> have False by linarith
   }
   note X = this
@@ -382,22 +382,22 @@
 lemma ord_inj' :
   assumes finite: "finite (carrier G)"
   assumes a: "a \<in> carrier G"
-  shows "inj_on (\<lambda> x . a (^) x) {1 .. ord a}"
+  shows "inj_on (\<lambda> x . a [^] x) {1 .. ord a}"
 proof (rule inj_onI, rule ccontr)
   fix x y :: nat
-  assume A:"x \<in> {1 .. ord a}" "y \<in> {1 .. ord a}" "a (^) x = a (^) y" "x\<noteq>y"
+  assume A:"x \<in> {1 .. ord a}" "y \<in> {1 .. ord a}" "a [^] x = a [^] y" "x\<noteq>y"
   { assume "x < ord a" "y < ord a"
     hence False using ord_inj[OF assms] A unfolding inj_on_def by fastforce
   }
   moreover
   { assume "x = ord a" "y < ord a"
-    hence "a (^) y = a (^) (0::nat)" using pow_ord_eq_1[OF assms] A by auto
+    hence "a [^] y = a [^] (0::nat)" using pow_ord_eq_1[OF assms] A by auto
     hence "y=0" using ord_inj[OF assms] \<open>y < ord a\<close> unfolding inj_on_def by force
     hence False using A by fastforce
   }
   moreover
   { assume "y = ord a" "x < ord a"
-    hence "a (^) x = a (^) (0::nat)" using pow_ord_eq_1[OF assms] A by auto
+    hence "a [^] x = a [^] (0::nat)" using pow_ord_eq_1[OF assms] A by auto
     hence "x=0" using ord_inj[OF assms] \<open>x < ord a\<close> unfolding inj_on_def by force
     hence False using A by fastforce
   }
@@ -406,35 +406,35 @@
 
 lemma ord_elems :
   assumes "finite (carrier G)" "a \<in> carrier G"
-  shows "{a(^)x | x. x \<in> (UNIV :: nat set)} = {a(^)x | x. x \<in> {0 .. ord a - 1}}" (is "?L = ?R")
+  shows "{a[^]x | x. x \<in> (UNIV :: nat set)} = {a[^]x | x. x \<in> {0 .. ord a - 1}}" (is "?L = ?R")
 proof
   show "?R \<subseteq> ?L" by blast
   { fix y assume "y \<in> ?L"
-    then obtain x::nat where x:"y = a(^)x" by auto
+    then obtain x::nat where x:"y = a[^]x" by auto
     define r where "r = x mod ord a"
     then obtain q where q:"x = q * ord a + r" using mod_eqD by atomize_elim presburger
-    hence "y = (a(^)ord a)(^)q \<otimes> a(^)r"
+    hence "y = (a[^]ord a)[^]q \<otimes> a[^]r"
       using x assms by (simp add: mult.commute nat_pow_mult nat_pow_pow)
-    hence "y = a(^)r" using assms by (simp add: pow_ord_eq_1)
+    hence "y = a[^]r" using assms by (simp add: pow_ord_eq_1)
     have "r < ord a" using ord_ge_1[OF assms] by (simp add: r_def)
     hence "r \<in> {0 .. ord a - 1}" by (force simp: r_def)
-    hence "y \<in> {a(^)x | x. x \<in> {0 .. ord a - 1}}" using \<open>y=a(^)r\<close> by blast
+    hence "y \<in> {a[^]x | x. x \<in> {0 .. ord a - 1}}" using \<open>y=a[^]r\<close> by blast
   }
   thus "?L \<subseteq> ?R" by auto
 qed
 
 lemma ord_dvd_pow_eq_1 :
-  assumes "finite (carrier G)" "a \<in> carrier G" "a (^) k = \<one>"
+  assumes "finite (carrier G)" "a \<in> carrier G" "a [^] k = \<one>"
   shows "ord a dvd k"
 proof -
   define r where "r = k mod ord a"
   then obtain q where q:"k = q*ord a + r" using mod_eqD by atomize_elim presburger
-  hence "a(^)k = (a(^)ord a)(^)q \<otimes> a(^)r"
+  hence "a[^]k = (a[^]ord a)[^]q \<otimes> a[^]r"
       using assms by (simp add: mult.commute nat_pow_mult nat_pow_pow)
-  hence "a(^)k = a(^)r" using assms by (simp add: pow_ord_eq_1)
-  hence "a(^)r = \<one>" using assms(3) by simp
+  hence "a[^]k = a[^]r" using assms by (simp add: pow_ord_eq_1)
+  hence "a[^]r = \<one>" using assms(3) by simp
   have "r < ord a" using ord_ge_1[OF assms(1-2)] by (simp add: r_def)
-  hence "r = 0" using \<open>a(^)r = \<one>\<close> ord_def[of a] ord_min[of r a] assms(1-2) by linarith
+  hence "r = 0" using \<open>a[^]r = \<one>\<close> ord_def[of a] ord_min[of r a] assms(1-2) by linarith
   thus ?thesis using q by simp
 qed
 
@@ -450,14 +450,14 @@
 lemma ord_pow_dvd_ord_elem :
   assumes finite[simp]: "finite (carrier G)"
   assumes a[simp]:"a \<in> carrier G"
-  shows "ord (a(^)n) = ord a div gcd n (ord a)"
+  shows "ord (a[^]n) = ord a div gcd n (ord a)"
 proof -
-  have "(a(^)n) (^) ord a = (a (^) ord a) (^) n"
+  have "(a[^]n) [^] ord a = (a [^] ord a) [^] n"
     by (simp add: mult.commute nat_pow_pow)
-  hence "(a(^)n) (^) ord a = \<one>" by (simp add: pow_ord_eq_1)
+  hence "(a[^]n) [^] ord a = \<one>" by (simp add: pow_ord_eq_1)
   obtain q where "n * (ord a div gcd n (ord a)) = ord a * q" by (rule dvd_gcd)
-  hence "(a(^)n) (^) (ord a div gcd n (ord a)) = (a (^) ord a)(^)q"  by (simp add : nat_pow_pow)
-  hence pow_eq_1: "(a(^)n) (^) (ord a div gcd n (ord a)) = \<one>"
+  hence "(a[^]n) [^] (ord a div gcd n (ord a)) = (a [^] ord a)[^]q"  by (simp add : nat_pow_pow)
+  hence pow_eq_1: "(a[^]n) [^] (ord a div gcd n (ord a)) = \<one>"
      by (auto simp add : pow_ord_eq_1[of a])
   have "ord a \<ge> 1" using ord_ge_1 by simp
   have ge_1:"ord a div gcd n (ord a) \<ge> 1"
@@ -471,12 +471,12 @@
     have "ord a div gcd n (ord a) \<le> ord a" by simp
     thus ?thesis using \<open>ord a \<le> order G\<close> by linarith
   qed
-  hence ord_gcd_elem:"ord a div gcd n (ord a) \<in> {d \<in> {1..order G}. (a(^)n) (^) d = \<one>}"
+  hence ord_gcd_elem:"ord a div gcd n (ord a) \<in> {d \<in> {1..order G}. (a[^]n) [^] d = \<one>}"
     using ge_1 pow_eq_1 by force
   { fix d :: nat
-    assume d_elem:"d \<in> {d \<in> {1..order G}. (a(^)n) (^) d = \<one>}"
+    assume d_elem:"d \<in> {d \<in> {1..order G}. (a[^]n) [^] d = \<one>}"
     assume d_lt:"d < ord a div gcd n (ord a)"
-    hence pow_nd:"a(^)(n*d)  = \<one>" using d_elem
+    hence pow_nd:"a[^](n*d)  = \<one>" using d_elem
       by (simp add : nat_pow_pow)
     hence "ord a dvd n*d" using assms by (auto simp add : ord_dvd_pow_eq_1)
     then obtain q where "ord a * q = n*d" by (metis dvd_mult_div_cancel)
@@ -500,9 +500,9 @@
     have "d > 0" using d_elem by simp
     hence "ord a div gcd n (ord a) \<le> d" using dvd_d by (simp add : Nat.dvd_imp_le)
     hence False using d_lt by simp
-  } hence ord_gcd_min: "\<And> d . d \<in> {d \<in> {1..order G}. (a(^)n) (^) d = \<one>}
+  } hence ord_gcd_min: "\<And> d . d \<in> {d \<in> {1..order G}. (a[^]n) [^] d = \<one>}
                         \<Longrightarrow> d\<ge>ord a div gcd n (ord a)" by fastforce
-  have fin:"finite {d \<in> {1..order G}. (a(^)n) (^) d = \<one>}" by auto
+  have fin:"finite {d \<in> {1..order G}. (a[^]n) [^] d = \<one>}" by auto
   thus ?thesis using Min_eqI[OF fin ord_gcd_min ord_gcd_elem]
     unfolding ord_def by simp
 qed
@@ -519,33 +519,33 @@
 lemma element_generates_subgroup:
   assumes finite[simp]: "finite (carrier G)"
   assumes a[simp]: "a \<in> carrier G"
-  shows "subgroup {a (^) i | i. i \<in> {0 .. ord a - 1}} G"
+  shows "subgroup {a [^] i | i. i \<in> {0 .. ord a - 1}} G"
 proof
-  show "{a(^)i | i. i \<in> {0 .. ord a - 1} } \<subseteq> carrier G" by auto
+  show "{a[^]i | i. i \<in> {0 .. ord a - 1} } \<subseteq> carrier G" by auto
 next
   fix x y
-  assume A: "x \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}" "y \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}"
-  obtain i::nat where i:"x = a(^)i" and i2:"i \<in> UNIV" using A by auto
-  obtain j::nat where j:"y = a(^)j" and j2:"j \<in> UNIV" using A by auto
-  have "a(^)(i+j) \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}" using ord_elems[OF assms] A by auto
-  thus "x \<otimes> y \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}"
+  assume A: "x \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}" "y \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}"
+  obtain i::nat where i:"x = a[^]i" and i2:"i \<in> UNIV" using A by auto
+  obtain j::nat where j:"y = a[^]j" and j2:"j \<in> UNIV" using A by auto
+  have "a[^](i+j) \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}" using ord_elems[OF assms] A by auto
+  thus "x \<otimes> y \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}"
     using i j a ord_elems assms by (auto simp add: nat_pow_mult)
 next
-  show "\<one> \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}" by force
+  show "\<one> \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}" by force
 next
-  fix x assume x: "x \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}"
+  fix x assume x: "x \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}"
   hence x_in_carrier: "x \<in> carrier G" by auto
-  then obtain d::nat where d:"x (^) d = \<one>" and "d\<ge>1"
+  then obtain d::nat where d:"x [^] d = \<one>" and "d\<ge>1"
     using finite_group_elem_finite_ord by auto
-  have inv_1:"x(^)(d - 1) \<otimes> x = \<one>" using \<open>d\<ge>1\<close> d nat_pow_Suc[of x "d - 1"] by simp
-  have elem:"x (^) (d - 1) \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}"
+  have inv_1:"x[^](d - 1) \<otimes> x = \<one>" using \<open>d\<ge>1\<close> d nat_pow_Suc[of x "d - 1"] by simp
+  have elem:"x [^] (d - 1) \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}"
   proof -
-    obtain i::nat where i:"x = a(^)i" using x by auto
-    hence "x(^)(d - 1) \<in> {a(^)i | i. i \<in> (UNIV::nat set)}" by (auto simp add: nat_pow_pow)
+    obtain i::nat where i:"x = a[^]i" using x by auto
+    hence "x[^](d - 1) \<in> {a[^]i | i. i \<in> (UNIV::nat set)}" by (auto simp add: nat_pow_pow)
     thus ?thesis using ord_elems[of a] by auto
   qed
-  have inv:"inv x = x(^)(d - 1)" using inv_equality[OF inv_1] x_in_carrier by blast
-  thus "inv x \<in> {a(^)i | i. i \<in> {0 .. ord a - 1}}" using elem inv by auto
+  have inv:"inv x = x[^](d - 1)" using inv_equality[OF inv_1] x_in_carrier by blast
+  thus "inv x \<in> {a[^]i | i. i \<in> {0 .. ord a - 1}}" using elem inv by auto
 qed
 
 lemma ord_dvd_group_order :
@@ -553,13 +553,13 @@
   assumes a[simp]: "a \<in> carrier G"
   shows "ord a dvd order G"
 proof -
-  have card_dvd:"card {a(^)i | i. i \<in> {0 .. ord a - 1}} dvd card (carrier G)"
+  have card_dvd:"card {a[^]i | i. i \<in> {0 .. ord a - 1}} dvd card (carrier G)"
     using lagrange_dvd element_generates_subgroup unfolding order_def by simp
-  have "inj_on (\<lambda> i . a(^)i) {0..ord a - 1}" using ord_inj by simp
-  hence cards_eq:"card ( (\<lambda> i . a(^)i) ` {0..ord a - 1}) = card {0..ord a - 1}"
-    using card_image[of "\<lambda> i . a(^)i" "{0..ord a - 1}"] by auto
-  have "(\<lambda> i . a(^)i) ` {0..ord a - 1} = {a(^)i | i. i \<in> {0..ord a - 1}}" by auto
-  hence "card {a(^)i | i. i \<in> {0..ord a - 1}} = card {0..ord a - 1}" using cards_eq by simp
+  have "inj_on (\<lambda> i . a[^]i) {0..ord a - 1}" using ord_inj by simp
+  hence cards_eq:"card ( (\<lambda> i . a[^]i) ` {0..ord a - 1}) = card {0..ord a - 1}"
+    using card_image[of "\<lambda> i . a[^]i" "{0..ord a - 1}"] by auto
+  have "(\<lambda> i . a[^]i) ` {0..ord a - 1} = {a[^]i | i. i \<in> {0..ord a - 1}}" by auto
+  hence "card {a[^]i | i. i \<in> {0..ord a - 1}} = card {0..ord a - 1}" using cards_eq by simp
   also have "\<dots> = ord a" using ord_ge_1[of a] by simp
   finally show ?thesis using card_dvd by (simp add: order_def)
 qed
@@ -580,7 +580,7 @@
 lemma mult_mult_of: "mult (mult_of R) = mult R"
  by (simp add: mult_of_def)
 
-lemma nat_pow_mult_of: "op (^)\<^bsub>mult_of R\<^esub> = (op (^)\<^bsub>R\<^esub> :: _ \<Rightarrow> nat \<Rightarrow> _)"
+lemma nat_pow_mult_of: "op [^]\<^bsub>mult_of R\<^esub> = (op [^]\<^bsub>R\<^esub> :: _ \<Rightarrow> nat \<Rightarrow> _)"
   by (simp add: mult_of_def fun_eq_iff nat_pow_def)
 
 lemma one_mult_of: "\<one>\<^bsub>mult_of R\<^esub> = \<one>\<^bsub>R\<^esub>"
@@ -609,7 +609,7 @@
 lemma (in monoid) Units_pow_closed :
   fixes d :: nat
   assumes "x \<in> Units G"
-  shows "x (^) d \<in> Units G"
+  shows "x [^] d \<in> Units G"
     by (metis assms group.is_monoid monoid.nat_pow_closed units_group units_of_carrier units_of_pow)
 
 lemma (in comm_monoid) is_monoid:
@@ -644,7 +644,7 @@
   have "\<And>x. eval R R id x f \<noteq> \<zero>"
   proof -
     fix x
-    have "(\<Oplus>i\<in>{..deg R f}. id (coeff P f i) \<otimes> x (^) i) \<noteq> \<zero>"
+    have "(\<Oplus>i\<in>{..deg R f}. id (coeff P f i) \<otimes> x [^] i) \<noteq> \<zero>"
       using 0 lcoeff_nonzero_nonzero[where p = f] by simp
     thus "eval R R id x f \<noteq> \<zero>" using 0 unfolding eval_def P_def by simp
   qed
@@ -697,7 +697,7 @@
   fixes p d :: nat
   assumes finite:"finite (carrier R)"
   assumes d_neq_zero : "d \<noteq> 0"
-  shows "card {x \<in> carrier R. x (^) d = \<one>} \<le> d"
+  shows "card {x \<in> carrier R. x [^] d = \<one>} \<le> d"
 proof -
   let ?f = "monom (UP R) \<one>\<^bsub>R\<^esub> d \<ominus>\<^bsub> (UP R)\<^esub> monom (UP R) \<one>\<^bsub>R\<^esub> 0"
   have one_in_carrier:"\<one> \<in> carrier R" by simp
@@ -708,9 +708,9 @@
   have roots_bound:"finite {a \<in> carrier R . eval R R id a ?f = \<zero>} \<and>
                     card {a \<in> carrier R . eval R R id a ?f = \<zero>} \<le> deg R ?f"
                     using finite by (intro R.roots_bound[OF _ f_not_zero]) simp
-  have subs:"{x \<in> carrier R. x (^) d = \<one>} \<subseteq> {a \<in> carrier R . eval R R id a ?f = \<zero>}"
+  have subs:"{x \<in> carrier R. x [^] d = \<one>} \<subseteq> {a \<in> carrier R . eval R R id a ?f = \<zero>}"
     by (auto simp: R.evalRR_simps)
-  then have "card {x \<in> carrier R. x (^) d = \<one>} \<le>
+  then have "card {x \<in> carrier R. x [^] d = \<one>} \<le>
         card {a \<in> carrier R. eval R R id a ?f = \<zero>}" using finite by (simp add : card_mono)
   thus ?thesis using \<open>deg R ?f = d\<close> roots_bound by linarith
 qed
@@ -728,7 +728,7 @@
 \<close>
 
 lemma (in group) pow_order_eq_1:
-  assumes "finite (carrier G)" "x \<in> carrier G" shows "x (^) order G = \<one>"
+  assumes "finite (carrier G)" "x \<in> carrier G" shows "x [^] order G = \<one>"
   using assms by (metis nat_pow_pow ord_dvd_group_order pow_ord_eq_1 dvdE nat_pow_one)
 
 (* XXX remove in AFP devel, replaced by div_eq_dividend_iff *)
@@ -742,7 +742,7 @@
 lemma (in group)
   assumes finite': "finite (carrier G)"
   assumes "a \<in> carrier G"
-  shows pow_ord_eq_ord_iff: "group.ord G (a (^) k) = ord a \<longleftrightarrow> coprime k (ord a)" (is "?L \<longleftrightarrow> ?R")
+  shows pow_ord_eq_ord_iff: "group.ord G (a [^] k) = ord a \<longleftrightarrow> coprime k (ord a)" (is "?L \<longleftrightarrow> ?R")
 proof
   assume A: ?L then show ?R
     using assms ord_ge_1 [OF assms]
@@ -762,32 +762,32 @@
   note mult_of_simps[simp]
   have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of)
 
-  interpret G:group "mult_of R" rewrites "op (^)\<^bsub>mult_of R\<^esub> = (op (^) :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>"
+  interpret G:group "mult_of R" rewrites "op [^]\<^bsub>mult_of R\<^esub> = (op [^] :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>"
     by (rule field_mult_group) simp_all
 
   from exists
   obtain a where a:"a \<in> carrier (mult_of R)" and ord_a: "group.ord (mult_of R) a = d"
     by (auto simp add: card_gt_0_iff)
 
-  have set_eq1:"{a(^)n| n. n \<in> {1 .. d}} = {x \<in> carrier (mult_of R). x (^) d = \<one>}"
+  have set_eq1:"{a[^]n| n. n \<in> {1 .. d}} = {x \<in> carrier (mult_of R). x [^] d = \<one>}"
   proof (rule card_seteq)
-    show "finite {x \<in> carrier (mult_of R). x (^) d = \<one>}" using finite by auto
+    show "finite {x \<in> carrier (mult_of R). x [^] d = \<one>}" using finite by auto
 
-    show "{a(^)n| n. n \<in> {1 ..d}} \<subseteq> {x \<in> carrier (mult_of R). x(^)d = \<one>}"
+    show "{a[^]n| n. n \<in> {1 ..d}} \<subseteq> {x \<in> carrier (mult_of R). x[^]d = \<one>}"
     proof
-      fix x assume "x \<in> {a(^)n | n. n \<in> {1 .. d}}"
-      then obtain n where n:"x = a(^)n \<and> n \<in> {1 .. d}" by auto
-      have "x(^)d =(a(^)d)(^)n" using n a ord_a by (simp add:nat_pow_pow mult.commute)
-      hence "x(^)d = \<one>" using ord_a G.pow_ord_eq_1[OF finite' a] by fastforce
-      thus "x \<in> {x \<in> carrier (mult_of R). x(^)d = \<one>}" using G.nat_pow_closed[OF a] n by blast
+      fix x assume "x \<in> {a[^]n | n. n \<in> {1 .. d}}"
+      then obtain n where n:"x = a[^]n \<and> n \<in> {1 .. d}" by auto
+      have "x[^]d =(a[^]d)[^]n" using n a ord_a by (simp add:nat_pow_pow mult.commute)
+      hence "x[^]d = \<one>" using ord_a G.pow_ord_eq_1[OF finite' a] by fastforce
+      thus "x \<in> {x \<in> carrier (mult_of R). x[^]d = \<one>}" using G.nat_pow_closed[OF a] n by blast
     qed
 
-    show "card {x \<in> carrier (mult_of R). x (^) d = \<one>} \<le> card {a(^)n | n. n \<in> {1 .. d}}"
+    show "card {x \<in> carrier (mult_of R). x [^] d = \<one>} \<le> card {a[^]n | n. n \<in> {1 .. d}}"
     proof -
-      have *:"{a(^)n | n. n \<in> {1 .. d }} = ((\<lambda> n. a(^)n) ` {1 .. d})" by auto
+      have *:"{a[^]n | n. n \<in> {1 .. d }} = ((\<lambda> n. a[^]n) ` {1 .. d})" by auto
       have "0 < order (mult_of R)" unfolding order_mult_of[OF finite]
         using card_mono[OF finite, of "{\<zero>, \<one>}"] by (simp add: order_def)
-      have "card {x \<in> carrier (mult_of R). x (^) d = \<one>} \<le> card {x \<in> carrier R. x (^) d = \<one>}"
+      have "card {x \<in> carrier (mult_of R). x [^] d = \<one>} \<le> card {x \<in> carrier R. x [^] d = \<one>}"
         using finite by (auto intro: card_mono)
       also have "\<dots> \<le> d" using \<open>0 < order (mult_of R)\<close> num_roots_le_deg[OF finite, of d]
         by (simp add : dvd_pos_nat[OF _ \<open>d dvd order (mult_of R)\<close>])
@@ -796,20 +796,20 @@
   qed
 
   have set_eq2:"{x \<in> carrier (mult_of R) . group.ord (mult_of R) x = d}
-                = (\<lambda> n . a(^)n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a(^)n) = d}" (is "?L = ?R")
+                = (\<lambda> n . a[^]n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d}" (is "?L = ?R")
   proof
     { fix x assume x:"x \<in> (carrier (mult_of R)) \<and> group.ord (mult_of R) x = d"
-      hence "x \<in> {x \<in> carrier (mult_of R). x (^) d = \<one>}"
+      hence "x \<in> {x \<in> carrier (mult_of R). x [^] d = \<one>}"
         by (simp add: G.pow_ord_eq_1[OF finite', of x, symmetric])
-      then obtain n where n:"x = a(^)n \<and> n \<in> {1 .. d}" using set_eq1 by blast
+      then obtain n where n:"x = a[^]n \<and> n \<in> {1 .. d}" using set_eq1 by blast
       hence "x \<in> ?R" using x by fast
     } thus "?L \<subseteq> ?R" by blast
     show "?R \<subseteq> ?L" using a by (auto simp add: carrier_mult_of[symmetric] simp del: carrier_mult_of)
   qed
-  have "inj_on (\<lambda> n . a(^)n) {n \<in> {1 .. d}. group.ord (mult_of R) (a(^)n) = d}"
+  have "inj_on (\<lambda> n . a[^]n) {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d}"
     using G.ord_inj'[OF finite' a, unfolded ord_a] unfolding inj_on_def by fast
-  hence "card ((\<lambda>n. a(^)n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a(^)n) = d})
-         = card {k \<in> {1 .. d}. group.ord (mult_of R) (a(^)k) = d}"
+  hence "card ((\<lambda>n. a[^]n) ` {n \<in> {1 .. d}. group.ord (mult_of R) (a[^]n) = d})
+         = card {k \<in> {1 .. d}. group.ord (mult_of R) (a[^]k) = d}"
          using card_image by blast
   thus ?thesis using set_eq2 G.pow_ord_eq_ord_iff[OF finite' \<open>a \<in> _\<close>, unfolded ord_a]
     by (simp add: phi'_def)
@@ -820,13 +820,13 @@
 
 theorem (in field) finite_field_mult_group_has_gen :
   assumes finite:"finite (carrier R)"
-  shows "\<exists> a \<in> carrier (mult_of R) . carrier (mult_of R) = {a(^)i | i::nat . i \<in> UNIV}"
+  shows "\<exists> a \<in> carrier (mult_of R) . carrier (mult_of R) = {a[^]i | i::nat . i \<in> UNIV}"
 proof -
   note mult_of_simps[simp]
   have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of)
 
   interpret G: group "mult_of R" rewrites
-      "op (^)\<^bsub>mult_of R\<^esub> = (op (^) :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>"
+      "op [^]\<^bsub>mult_of R\<^esub> = (op [^] :: _ \<Rightarrow> nat \<Rightarrow> _)" and "\<one>\<^bsub>mult_of R\<^esub> = \<one>"
     by (rule field_mult_group) (simp_all add: fun_eq_iff nat_pow_def)
 
   let ?N = "\<lambda> x . card {a \<in> carrier (mult_of R). group.ord (mult_of R) a  = x}"
@@ -888,17 +888,17 @@
   hence "?N (order (mult_of R)) > 0" using * by (simp add: phi'_nonzero)
   then obtain a where a:"a \<in> carrier (mult_of R)" and a_ord:"group.ord (mult_of R) a = order (mult_of R)"
     by (auto simp add: card_gt_0_iff)
-  hence set_eq:"{a(^)i | i::nat. i \<in> UNIV} = (\<lambda>x. a(^)x) ` {0 .. group.ord (mult_of R) a - 1}"
+  hence set_eq:"{a[^]i | i::nat. i \<in> UNIV} = (\<lambda>x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}"
     using G.ord_elems[OF finite'] by auto
-  have card_eq:"card ((\<lambda>x. a(^)x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 .. group.ord (mult_of R) a - 1}"
+  have card_eq:"card ((\<lambda>x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 .. group.ord (mult_of R) a - 1}"
     by (intro card_image G.ord_inj finite' a)
-  hence "card ((\<lambda> x . a(^)x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 ..order (mult_of R) - 1}"
+  hence "card ((\<lambda> x . a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 ..order (mult_of R) - 1}"
     using assms by (simp add: card_eq a_ord)
-  hence card_R_minus_1:"card {a(^)i | i::nat. i \<in> UNIV} =  order (mult_of R)"
+  hence card_R_minus_1:"card {a[^]i | i::nat. i \<in> UNIV} =  order (mult_of R)"
     using * by (subst set_eq) auto
-  have **:"{a(^)i | i::nat. i \<in> UNIV} \<subseteq> carrier (mult_of R)"
+  have **:"{a[^]i | i::nat. i \<in> UNIV} \<subseteq> carrier (mult_of R)"
     using G.nat_pow_closed[OF a] by auto
-  with _ have "carrier (mult_of R) = {a(^)i|i::nat. i \<in> UNIV}"
+  with _ have "carrier (mult_of R) = {a[^]i|i::nat. i \<in> UNIV}"
     by (rule card_seteq[symmetric]) (simp_all add: card_R_minus_1 finite order_def del: UNIV_I)
   thus ?thesis using a by blast
 qed
--- a/src/HOL/Algebra/Ring.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/Ring.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -146,13 +146,13 @@
 
 lemmas finsum_reindex = add.finprod_reindex
 
-(* The following would be wrong.  Needed is the equivalent of (^) for addition,
+(* The following would be wrong.  Needed is the equivalent of [^] for addition,
   or indeed the canonical embedding from Nat into the monoid.
 
 lemma finsum_const:
   assumes fin [simp]: "finite A"
       and a [simp]: "a : carrier G"
-    shows "finsum G (%x. a) A = a (^) card A"
+    shows "finsum G (%x. a) A = a [^] card A"
   using fin apply induct
   apply force
   apply (subst finsum_insert)
@@ -427,7 +427,7 @@
   a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus
 
 lemma (in semiring) nat_pow_zero:
-  "(n::nat) \<noteq> 0 \<Longrightarrow> \<zero> (^) n = \<zero>"
+  "(n::nat) \<noteq> 0 \<Longrightarrow> \<zero> [^] n = \<zero>"
   by (induct n) simp_all
 
 context semiring begin
--- a/src/HOL/Algebra/UnivPoly.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Algebra/UnivPoly.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -1175,7 +1175,7 @@
   eval :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme,
            'a => 'b, 'b, nat => 'a] => 'b"
   where "eval R S phi s = (\<lambda>p \<in> carrier (UP R).
-    \<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+    \<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
 
 context UP
 begin
@@ -1183,7 +1183,7 @@
 lemma eval_on_carrier:
   fixes S (structure)
   shows "p \<in> carrier P ==>
-  eval R S phi s p = (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. phi (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+  eval R S phi s p = (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. phi (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
   by (unfold eval_def, fold P_def) simp
 
 lemma eval_extensional:
@@ -1227,35 +1227,35 @@
   then show "eval R S h s (p \<oplus>\<^bsub>P\<^esub> q) = eval R S h s p \<oplus>\<^bsub>S\<^esub> eval R S h s q"
   proof (simp only: eval_on_carrier P.a_closed)
     from S R have
-      "(\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
+      "(\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) =
       (\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<oplus>\<^bsub>P\<^esub> q)<..max (deg R p) (deg R q)}.
-        h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+        h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp cong: S.finsum_cong
         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def del: coeff_add)
     also from R have "... =
         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..max (deg R p) (deg R q)}.
-          h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+          h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp add: ivl_disj_un_one)
     also from R S have "... =
-      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
-      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
+      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp cong: S.finsum_cong
         add: S.l_distr deg_aboveD ivl_disj_int_one Pi_def)
     also have "... =
         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p} \<union> {deg R p<..max (deg R p) (deg R q)}.
-          h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
+          h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
         (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q} \<union> {deg R q<..max (deg R p) (deg R q)}.
-          h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+          h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp only: ivl_disj_un_one max.cobounded1 max.cobounded2)
     also from R S have "... =
-      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
-      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
+      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp cong: S.finsum_cong
         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
     finally show
-      "(\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
-      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
-      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
+      "(\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) =
+      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
+      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" .
   qed
 next
   show "eval R S h s \<one>\<^bsub>P\<^esub> = \<one>\<^bsub>S\<^esub>"
@@ -1266,31 +1266,31 @@
   then show "eval R S h s (p \<otimes>\<^bsub>P\<^esub> q) = eval R S h s p \<otimes>\<^bsub>S\<^esub> eval R S h s q"
   proof (simp only: eval_on_carrier UP_mult_closed)
     from R S have
-      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
+      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) =
       (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<otimes>\<^bsub>P\<^esub> q)<..deg R p + deg R q}.
-        h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+        h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp cong: S.finsum_cong
         add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def
         del: coeff_mult)
     also from R have "... =
-      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp only: ivl_disj_un_one deg_mult_ring)
     also from R S have "... =
       (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}.
          \<Oplus>\<^bsub>S\<^esub> k \<in> {..i}.
            h (coeff P p k) \<otimes>\<^bsub>S\<^esub> h (coeff P q (i - k)) \<otimes>\<^bsub>S\<^esub>
-           (s (^)\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> (i - k)))"
+           (s [^]\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> (i - k)))"
       by (simp cong: S.finsum_cong add: S.nat_pow_mult Pi_def
         S.m_ac S.finsum_rdistr)
     also from R S have "... =
-      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
-      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
+      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
       by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac
         Pi_def)
     finally show
-      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
-      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
-      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
+      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) =
+      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
+      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" .
   qed
 qed
 
@@ -1314,13 +1314,13 @@
   shows "eval R S h s (monom P \<one> 1) = s"
 proof (simp only: eval_on_carrier monom_closed R.one_closed)
    from S have
-    "(\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
+    "(\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) =
     (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)} \<union> {deg R (monom P \<one> 1)<..1}.
-      h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+      h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
     by (simp cong: S.finsum_cong del: coeff_monom
       add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
   also have "... =
-    (\<Oplus>\<^bsub>S\<^esub> i \<in> {..1}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
+    (\<Oplus>\<^bsub>S\<^esub> i \<in> {..1}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)"
     by (simp only: ivl_disj_un_one deg_monom_le R.one_closed)
   also have "... = s"
   proof (cases "s = \<zero>\<^bsub>S\<^esub>")
@@ -1329,7 +1329,7 @@
     case False then show ?thesis by (simp add: S Pi_def)
   qed
   finally show "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (monom P \<one> 1)}.
-    h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) = s" .
+    h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = s" .
 qed
 
 end
@@ -1342,7 +1342,7 @@
 
 lemma (in UP_cring) monom_pow:
   assumes R: "a \<in> carrier R"
-  shows "(monom P a n) (^)\<^bsub>P\<^esub> m = monom P (a (^) m) (n * m)"
+  shows "(monom P a n) [^]\<^bsub>P\<^esub> m = monom P (a [^] m) (n * m)"
 proof (induct m)
   case 0 from R show ?case by simp
 next
@@ -1351,25 +1351,25 @@
 qed
 
 lemma (in ring_hom_cring) hom_pow [simp]:
-  "x \<in> carrier R ==> h (x (^) n) = h x (^)\<^bsub>S\<^esub> (n::nat)"
+  "x \<in> carrier R ==> h (x [^] n) = h x [^]\<^bsub>S\<^esub> (n::nat)"
   by (induct n) simp_all
 
 lemma (in UP_univ_prop) Eval_monom:
-  "r \<in> carrier R ==> Eval (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
+  "r \<in> carrier R ==> Eval (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n"
 proof -
   assume R: "r \<in> carrier R"
-  from R have "Eval (monom P r n) = Eval (monom P r 0 \<otimes>\<^bsub>P\<^esub> (monom P \<one> 1) (^)\<^bsub>P\<^esub> n)"
+  from R have "Eval (monom P r n) = Eval (monom P r 0 \<otimes>\<^bsub>P\<^esub> (monom P \<one> 1) [^]\<^bsub>P\<^esub> n)"
     by (simp del: monom_mult add: monom_mult [THEN sym] monom_pow)
   also
   from R eval_monom1 [where s = s, folded Eval_def]
-  have "... = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
+  have "... = h r \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n"
     by (simp add: eval_const [where s = s, folded Eval_def])
   finally show ?thesis .
 qed
 
 lemma (in UP_pre_univ_prop) eval_monom:
   assumes R: "r \<in> carrier R" and S: "s \<in> carrier S"
-  shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
+  shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n"
 proof -
   interpret UP_univ_prop R S h P s "eval R S h s"
     using UP_pre_univ_prop_axioms P_def R S
@@ -1411,11 +1411,11 @@
   interpret ring_hom_cring P S Phi by fact
   interpret ring_hom_cring P S Psi by fact
   have "Phi p =
-      Phi (\<Oplus>\<^bsub>P \<^esub>i \<in> {..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
+      Phi (\<Oplus>\<^bsub>P \<^esub>i \<in> {..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 [^]\<^bsub>P\<^esub> i)"
     by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
   also
   have "... =
-      Psi (\<Oplus>\<^bsub>P \<^esub>i\<in>{..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
+      Psi (\<Oplus>\<^bsub>P \<^esub>i\<in>{..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 [^]\<^bsub>P\<^esub> i)"
     by (simp add: Phi Psi P Pi_def comp_def)
   also have "... = Psi p"
     by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
@@ -1445,7 +1445,7 @@
 context monoid
 begin
 
-lemma nat_pow_eone[simp]: assumes x_in_G: "x \<in> carrier G" shows "x (^) (1::nat) = x"
+lemma nat_pow_eone[simp]: assumes x_in_G: "x \<in> carrier G" shows "x [^] (1::nat) = x"
   using nat_pow_Suc [of x 0] unfolding nat_pow_0 [of x] unfolding l_one [OF x_in_G] by simp
 
 end
@@ -1550,14 +1550,14 @@
 lemma long_div_theorem:
   assumes g_in_P [simp]: "g \<in> carrier P" and f_in_P [simp]: "f \<in> carrier P"
   and g_not_zero: "g \<noteq> \<zero>\<^bsub>P\<^esub>"
-  shows "\<exists>q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g)"
+  shows "\<exists>q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)[^]\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g)"
   using f_in_P
 proof (induct "deg R f" arbitrary: "f" rule: nat_less_induct)
   case (1 f)
   note f_in_P [simp] = "1.prems"
   let ?pred = "(\<lambda> q r (k::nat).
     (q \<in> carrier P) \<and> (r \<in> carrier P)
-    \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g))"
+    \<and> (lcoeff g)[^]\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g))"
   let ?lg = "lcoeff g" and ?lf = "lcoeff f"
   show ?case
   proof (cases "deg R f < deg R g")
@@ -1585,7 +1585,7 @@
       next
         case False note deg_f_nzero = False
         {
-          have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?f1"
+          have exist: "lcoeff g [^] ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?f1"
             by (simp add: minus_add r_neg sym [
               OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]])
           have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?f1) < deg R f"
@@ -1611,28 +1611,28 @@
             qed (simp_all add: deg_f_nzero)
           qed
           then obtain q' r' k'
-            where rem_desc: "?lg (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?f1) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
+            where rem_desc: "?lg [^] (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?f1) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
             and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
             and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
             using "1.hyps" using f1_in_carrier by blast
           show ?thesis
-          proof (rule exI3 [of _ "((?lg (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
-            show "(?lg (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((?lg (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
+          proof (rule exI3 [of _ "((?lg [^] k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
+            show "(?lg [^] (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((?lg [^] k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
             proof -
-              have "(?lg (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (?lg (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?f1)"
+              have "(?lg [^] (Suc k')) \<odot>\<^bsub>P\<^esub> f = (?lg [^] k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?f1)"
                 using smult_assoc1 [OF _ _ f_in_P] using exist by simp
-              also have "\<dots> = (?lg (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((?lg (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?f1))"
+              also have "\<dots> = (?lg [^] k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((?lg [^] k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?f1))"
                 using UP_smult_r_distr by simp
-              also have "\<dots> = (?lg (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
+              also have "\<dots> = (?lg [^] k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
                 unfolding rem_desc ..
-              also have "\<dots> = (?lg (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
-                using sym [OF a_assoc [of "?lg (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
+              also have "\<dots> = (?lg [^] k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
+                using sym [OF a_assoc [of "?lg [^] k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
                 using r'_in_carrier q'_in_carrier by simp
-              also have "\<dots> = (?lg (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+              also have "\<dots> = (?lg [^] k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
                 using q'_in_carrier by (auto simp add: m_comm)
-              also have "\<dots> = (((?lg (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+              also have "\<dots> = (((?lg [^] k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
                 using smult_assoc2 q'_in_carrier "1.prems" by auto
-              also have "\<dots> = ((?lg (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
+              also have "\<dots> = ((?lg [^] k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
                 using sym [OF l_distr] and q'_in_carrier by auto
               finally show ?thesis using m_comm q'_in_carrier by auto
             qed
@@ -1740,7 +1740,7 @@
   from deg_minus_monom [OF a R_not_trivial]
   have deg_g_nzero: "deg R ?g \<noteq> 0" by simp
   have "\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and>
-    lcoeff ?g (^) k \<odot>\<^bsub>P\<^esub> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R ?g)"
+    lcoeff ?g [^] k \<odot>\<^bsub>P\<^esub> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R ?g)"
     using long_div_theorem [OF _ f deg_nzero_nzero [OF deg_g_nzero]] a
     by auto
   then show ?thesis
--- a/src/HOL/Decision_Procs/Algebra_Aux.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Decision_Procs/Algebra_Aux.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -190,7 +190,7 @@
   with assms show "x \<ominus> y = \<zero>" by (simp add: minus_eq r_neg)
 qed
 
-lemma power2_eq_square: "x \<in> carrier R \<Longrightarrow> x (^) (2::nat) = x \<otimes> x"
+lemma power2_eq_square: "x \<in> carrier R \<Longrightarrow> x [^] (2::nat) = x \<otimes> x"
   by (simp add: numeral_eq_Suc)
 
 lemma eq_neg_iff_add_eq_0:
@@ -230,7 +230,7 @@
 
 end
 
-lemma (in cring) of_int_power [simp]: "\<guillemotleft>i ^ n\<guillemotright> = \<guillemotleft>i\<guillemotright> (^) n"
+lemma (in cring) of_int_power [simp]: "\<guillemotleft>i ^ n\<guillemotright> = \<guillemotleft>i\<guillemotright> [^] n"
   by (induct n) (simp_all add: m_ac)
 
 definition cring_class_ops :: "'a::comm_ring_1 ring"
@@ -267,7 +267,7 @@
 lemma minus_class: "x \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
   by (simp add: a_minus_def carrier_class plus_class uminus_class)
 
-lemma power_class: "x (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n"
+lemma power_class: "x [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
   by (induct n) (simp_all add: one_class times_class
     monoid.nat_pow_0 [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]]
     monoid.nat_pow_Suc [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]])
@@ -288,14 +288,14 @@
     and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
     and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
     and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
-    and "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n"
+    and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
     and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
     and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
     and "(Int.of_int (numeral m)::'a) = numeral m"
   by (simp_all add: cring_class class_simps)
 
 lemma (in domain) nat_pow_eq_0_iff [simp]:
-  "a \<in> carrier R \<Longrightarrow> (a (^) (n::nat) = \<zero>) = (a = \<zero> \<and> n \<noteq> 0)"
+  "a \<in> carrier R \<Longrightarrow> (a [^] (n::nat) = \<zero>) = (a = \<zero> \<and> n \<noteq> 0)"
   by (induct n) (auto simp add: integral_iff)
 
 lemma (in domain) square_eq_iff:
@@ -446,7 +446,7 @@
 
 lemma nonzero_power_divide:
   "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow>
-    (a \<oslash> b) (^) (n::nat) = a (^) n \<oslash> b (^) n"
+    (a \<oslash> b) [^] (n::nat) = a [^] n \<oslash> b [^] n"
   by (induct n) (simp_all add: nonzero_divide_divide_eq_left)
 
 lemma r_diff_distr:
@@ -504,7 +504,7 @@
     and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
     and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
     and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
-    and "(x::'a) (^)\<^bsub>cring_class_ops\<^esub> n = x ^ n"
+    and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
     and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
     and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
     and "(x::'a) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y"
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -58,7 +58,7 @@
   where
     "Ipol l (Pc c) = \<guillemotleft>c\<guillemotright>"
   | "Ipol l (Pinj i P) = Ipol (drop i l) P"
-  | "Ipol l (PX P x Q) = Ipol l P \<otimes> head l (^) x \<oplus> Ipol (drop 1 l) Q"
+  | "Ipol l (PX P x Q) = Ipol l P \<otimes> head l [^] x \<oplus> Ipol (drop 1 l) Q"
 
 lemma Ipol_Pc:
   "Ipol l (Pc 0) = \<zero>"
@@ -77,7 +77,7 @@
   | "Ipolex l (Add P Q) = Ipolex l P \<oplus> Ipolex l Q"
   | "Ipolex l (Sub P Q) = Ipolex l P \<ominus> Ipolex l Q"
   | "Ipolex l (Mul P Q) = Ipolex l P \<otimes> Ipolex l Q"
-  | "Ipolex l (Pow p n) = Ipolex l p (^) n"
+  | "Ipolex l (Pow p n) = Ipolex l p [^] n"
   | "Ipolex l (Neg P) = \<ominus> Ipolex l P"
 
 lemma Ipolex_Const:
@@ -302,7 +302,7 @@
        a_ac m_ac nat_pow_mult [symmetric] of_int_2)
 
 text \<open>Power\<close>
-lemma pow_ci: "in_carrier ls \<Longrightarrow> Ipol ls (pow n P) = Ipol ls P (^) n"
+lemma pow_ci: "in_carrier ls \<Longrightarrow> Ipol ls (pow n P) = Ipol ls P [^] n"
 proof (induct n arbitrary: P rule: less_induct)
   case (less k)
   consider "k = 0" | "k > 0" by arith
@@ -313,7 +313,7 @@
   next
     case 2
     then have "k div 2 < k" by arith
-    with less have *: "Ipol ls (pow (k div 2) (sqr P)) = Ipol ls (sqr P) (^) (k div 2)"
+    with less have *: "Ipol ls (pow (k div 2) (sqr P)) = Ipol ls (sqr P) [^] (k div 2)"
       by simp
     show ?thesis
     proof (cases "even k")
@@ -358,7 +358,7 @@
   where
     "Imon l (Mc c) = \<guillemotleft>c\<guillemotright>"
   | "Imon l (Minj i M) = Imon (drop i l) M"
-  | "Imon l (MX x M) = Imon (drop 1 l) M \<otimes> head l (^) x"
+  | "Imon l (MX x M) = Imon (drop 1 l) M \<otimes> head l [^] x"
 
 lemma (in cring) Imon_closed [simp]: "in_carrier l \<Longrightarrow> Imon l m \<in> carrier R"
   by (induct m arbitrary: l) simp_all
@@ -385,7 +385,7 @@
 lemma (in cring) Minj_pred_correct: "0 < i \<Longrightarrow> Imon (drop 1 l) (Minj_pred i M) = Imon l (Minj i M)"
   by (simp add: Minj_pred_def mkMinj_correct)
 
-lemma (in cring) mkMX_correct: "in_carrier l \<Longrightarrow> Imon l (mkMX i M) = Imon l M \<otimes> head l (^) i"
+lemma (in cring) mkMX_correct: "in_carrier l \<Longrightarrow> Imon l (mkMX i M) = Imon l M \<otimes> head l [^] i"
   by (induct M)
     (simp_all add: Minj_pred_correct [simplified] nat_pow_mult [symmetric] m_ac split: mon.split)
 
@@ -485,18 +485,18 @@
 next
   case (PX_MX P i Q j M)
   from \<open>in_carrier l\<close>
-  have eq1: "(Imon (drop (Suc 0) l) M \<otimes> head l (^) (j - i)) \<otimes>
-    Ipol l (snd (mfactor P (MX (j - i) M))) \<otimes> head l (^) i =
+  have eq1: "(Imon (drop (Suc 0) l) M \<otimes> head l [^] (j - i)) \<otimes>
+    Ipol l (snd (mfactor P (MX (j - i) M))) \<otimes> head l [^] i =
     Imon (drop (Suc 0) l) M \<otimes>
     Ipol l (snd (mfactor P (MX (j - i) M))) \<otimes>
-    (head l (^) (j - i) \<otimes> head l (^) i)"
+    (head l [^] (j - i) \<otimes> head l [^] i)"
     by (simp add: m_ac)
   from \<open>in_carrier l\<close>
-  have eq2: "(Imon (drop (Suc 0) l) M \<otimes> head l (^) j) \<otimes>
-    (Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \<otimes> head l (^) (i - j)) =
+  have eq2: "(Imon (drop (Suc 0) l) M \<otimes> head l [^] j) \<otimes>
+    (Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \<otimes> head l [^] (i - j)) =
     Imon (drop (Suc 0) l) M \<otimes>
     Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \<otimes>
-    (head l (^) (i - j) \<otimes> head l (^) j)"
+    (head l [^] (i - j) \<otimes> head l [^] j)"
     by (simp add: m_ac)
   from PX_MX \<open>in_carrier l\<close> show ?case
     by (simp add: mkPX_ci mkMinj_correct l_distr eq1 eq2 split_beta nat_pow_mult)
--- a/src/HOL/Decision_Procs/Reflective_Field.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Decision_Procs/Reflective_Field.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -39,7 +39,7 @@
   | "feval xs (FMul a b) = feval xs a \<otimes> feval xs b"
   | "feval xs (FNeg a) = \<ominus> feval xs a"
   | "feval xs (FDiv a b) = feval xs a \<oslash> feval xs b"
-  | "feval xs (FPow a n) = feval xs a (^) n"
+  | "feval xs (FPow a n) = feval xs a [^] n"
 
 lemma (in field) feval_Cnst:
   "feval xs (FCnst 0) = \<zero>"
@@ -96,7 +96,7 @@
   | "peval xs (PExpr1 (PSub a b)) = peval xs a \<ominus> peval xs b"
   | "peval xs (PExpr1 (PNeg a)) = \<ominus> peval xs a"
   | "peval xs (PExpr2 (PMul a b)) = peval xs a \<otimes> peval xs b"
-  | "peval xs (PExpr2 (PPow a n)) = peval xs a (^) n"
+  | "peval xs (PExpr2 (PPow a n)) = peval xs a [^] n"
 
 lemma (in field) peval_Cnst:
   "peval xs (PExpr1 (PCnst 0)) = \<zero>"
@@ -264,7 +264,7 @@
 
 lemma (in field) isin_correct':
   "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow>
-    peval xs e' = peval xs e (^) (n - p) \<otimes> peval xs e''"
+    peval xs e' = peval xs e [^] (n - p) \<otimes> peval xs e''"
   "in_carrier xs \<Longrightarrow> isin e n e' 1 = Some (p, e'') \<Longrightarrow> p \<le> n"
   using isin_correct [where m = 1] by simp_all
 
--- a/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -13,8 +13,8 @@
 
 lemma (in cring)
   assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R"
-  shows "\<guillemotleft>4\<guillemotright> \<otimes> x (^) (5::nat) \<otimes> y (^) (3::nat) \<otimes> x (^) (2::nat) \<otimes> \<guillemotleft>3\<guillemotright> \<oplus> x \<otimes> z \<oplus> \<guillemotleft>3\<guillemotright> (^) (5::nat) =
-    \<guillemotleft>12\<guillemotright> \<otimes> x (^) (7::nat) \<otimes> y (^) (3::nat) \<oplus> z \<otimes> x \<oplus> \<guillemotleft>243\<guillemotright>"
+  shows "\<guillemotleft>4\<guillemotright> \<otimes> x [^] (5::nat) \<otimes> y [^] (3::nat) \<otimes> x [^] (2::nat) \<otimes> \<guillemotleft>3\<guillemotright> \<oplus> x \<otimes> z \<oplus> \<guillemotleft>3\<guillemotright> [^] (5::nat) =
+    \<guillemotleft>12\<guillemotright> \<otimes> x [^] (7::nat) \<otimes> y [^] (3::nat) \<oplus> z \<otimes> x \<oplus> \<guillemotleft>243\<guillemotright>"
   by ring
 
 lemma "((x::int) + y) ^ 2  = x ^ 2 + y ^ 2 + 2 * x * y"
@@ -22,7 +22,7 @@
 
 lemma (in cring)
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x \<oplus> y) (^) (2::nat)  = x (^) (2::nat) \<oplus> y (^) (2::nat) \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> x \<otimes> y"
+  shows "(x \<oplus> y) [^] (2::nat)  = x [^] (2::nat) \<oplus> y [^] (2::nat) \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> x \<otimes> y"
   by ring
 
 lemma "((x::int) + y) ^ 3 = x ^ 3 + y ^ 3 + 3 * x ^ 2 * y + 3 * y ^ 2 * x"
@@ -30,8 +30,8 @@
 
 lemma (in cring)
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x \<oplus> y) (^) (3::nat) =
-    x (^) (3::nat) \<oplus> y (^) (3::nat) \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> x (^) (2::nat) \<otimes> y \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> y (^) (2::nat) \<otimes> x"
+  shows "(x \<oplus> y) [^] (3::nat) =
+    x [^] (3::nat) \<oplus> y [^] (3::nat) \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> x [^] (2::nat) \<otimes> y \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> y [^] (2::nat) \<otimes> x"
   by ring
 
 lemma "((x::int) - y) ^ 3 = x ^ 3 + 3 * x * y ^ 2 + (- 3) * y * x ^ 2 - y ^ 3"
@@ -39,8 +39,8 @@
 
 lemma (in cring)
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x \<ominus> y) (^) (3::nat) =
-    x (^) (3::nat) \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> x \<otimes> y (^) (2::nat) \<oplus> (\<ominus> \<guillemotleft>3\<guillemotright>) \<otimes> y \<otimes> x (^) (2::nat) \<ominus> y (^) (3::nat)"
+  shows "(x \<ominus> y) [^] (3::nat) =
+    x [^] (3::nat) \<oplus> \<guillemotleft>3\<guillemotright> \<otimes> x \<otimes> y [^] (2::nat) \<oplus> (\<ominus> \<guillemotleft>3\<guillemotright>) \<otimes> y \<otimes> x [^] (2::nat) \<ominus> y [^] (3::nat)"
   by ring
 
 lemma "((x::int) - y) ^ 2 = x ^ 2 + y ^ 2 - 2 * x * y"
@@ -48,7 +48,7 @@
 
 lemma (in cring)
   assumes "x \<in> carrier R" "y \<in> carrier R"
-  shows "(x \<ominus> y) (^) (2::nat) = x (^) (2::nat) \<oplus> y (^) (2::nat) \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> x \<otimes> y"
+  shows "(x \<ominus> y) [^] (2::nat) = x [^] (2::nat) \<oplus> y [^] (2::nat) \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> x \<otimes> y"
   by ring
 
 lemma " ((a::int) + b + c) ^ 2 = a ^ 2 + b ^ 2 + c ^ 2 + 2 * a * b + 2 * b * c + 2 * a * c"
@@ -56,8 +56,8 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
-  shows " (a \<oplus> b \<oplus> c) (^) (2::nat) =
-    a (^) (2::nat) \<oplus> b (^) (2::nat) \<oplus> c (^) (2::nat) \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> b \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> b \<otimes> c \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> c"
+  shows " (a \<oplus> b \<oplus> c) [^] (2::nat) =
+    a [^] (2::nat) \<oplus> b [^] (2::nat) \<oplus> c [^] (2::nat) \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> b \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> b \<otimes> c \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> c"
   by ring
 
 lemma "((a::int) - b - c) ^ 2 = a ^ 2 + b ^ 2 + c ^ 2 - 2 * a * b + 2 * b * c - 2 * a * c"
@@ -65,8 +65,8 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
-  shows "(a \<ominus> b \<ominus> c) (^) (2::nat) =
-    a (^) (2::nat) \<oplus> b (^) (2::nat) \<oplus> c (^) (2::nat) \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> b \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> b \<otimes> c \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> c"
+  shows "(a \<ominus> b \<ominus> c) [^] (2::nat) =
+    a [^] (2::nat) \<oplus> b [^] (2::nat) \<oplus> c [^] (2::nat) \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> b \<oplus> \<guillemotleft>2\<guillemotright> \<otimes> b \<otimes> c \<ominus> \<guillemotleft>2\<guillemotright> \<otimes> a \<otimes> c"
   by ring
 
 lemma "(a::int) * b + a * c = a * (b + c)"
@@ -82,7 +82,7 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R"
-  shows "a (^) (2::nat) \<ominus> b (^) (2::nat) = (a \<ominus> b) \<otimes> (a \<oplus> b)"
+  shows "a [^] (2::nat) \<ominus> b [^] (2::nat) = (a \<ominus> b) \<otimes> (a \<oplus> b)"
   by ring
 
 lemma "(a::int) ^ 3 - b ^ 3 = (a - b) * (a ^ 2 + a * b + b ^ 2)"
@@ -90,7 +90,7 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R"
-  shows "a (^) (3::nat) \<ominus> b (^) (3::nat) = (a \<ominus> b) \<otimes> (a (^) (2::nat) \<oplus> a \<otimes> b \<oplus> b (^) (2::nat))"
+  shows "a [^] (3::nat) \<ominus> b [^] (3::nat) = (a \<ominus> b) \<otimes> (a [^] (2::nat) \<oplus> a \<otimes> b \<oplus> b [^] (2::nat))"
   by ring
 
 lemma "(a::int) ^ 3 + b ^ 3 = (a + b) * (a ^ 2 - a * b + b ^ 2)"
@@ -98,7 +98,7 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R"
-  shows "a (^) (3::nat) \<oplus> b (^) (3::nat) = (a \<oplus> b) \<otimes> (a (^) (2::nat) \<ominus> a \<otimes> b \<oplus> b (^) (2::nat))"
+  shows "a [^] (3::nat) \<oplus> b [^] (3::nat) = (a \<oplus> b) \<otimes> (a [^] (2::nat) \<ominus> a \<otimes> b \<oplus> b [^] (2::nat))"
   by ring
 
 lemma "(a::int) ^ 4 - b ^ 4 = (a - b) * (a + b) * (a ^ 2 + b ^ 2)"
@@ -106,7 +106,7 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R"
-  shows "a (^) (4::nat) \<ominus> b (^) (4::nat) = (a \<ominus> b) \<otimes> (a \<oplus> b) \<otimes> (a (^) (2::nat) \<oplus> b (^) (2::nat))"
+  shows "a [^] (4::nat) \<ominus> b [^] (4::nat) = (a \<ominus> b) \<otimes> (a \<oplus> b) \<otimes> (a [^] (2::nat) \<oplus> b [^] (2::nat))"
   by ring
 
 lemma "(a::int) ^ 10 - b ^ 10 =
@@ -116,11 +116,11 @@
 
 lemma (in cring)
   assumes "a \<in> carrier R" "b \<in> carrier R"
-  shows "a (^) (10::nat) \<ominus> b (^) (10::nat) =
-  (a \<ominus> b) \<otimes> (a (^) (9::nat) \<oplus> a (^) (8::nat) \<otimes> b \<oplus> a (^) (7::nat) \<otimes> b (^) (2::nat) \<oplus>
-    a (^) (6::nat) \<otimes> b (^) (3::nat) \<oplus> a (^) (5::nat) \<otimes> b (^) (4::nat) \<oplus>
-    a (^) (4::nat) \<otimes> b (^) (5::nat) \<oplus> a (^) (3::nat) \<otimes> b (^) (6::nat) \<oplus>
-    a (^) (2::nat) \<otimes> b (^) (7::nat) \<oplus> a \<otimes> b (^) (8::nat) \<oplus> b (^) (9::nat))"
+  shows "a [^] (10::nat) \<ominus> b [^] (10::nat) =
+  (a \<ominus> b) \<otimes> (a [^] (9::nat) \<oplus> a [^] (8::nat) \<otimes> b \<oplus> a [^] (7::nat) \<otimes> b [^] (2::nat) \<oplus>
+    a [^] (6::nat) \<otimes> b [^] (3::nat) \<oplus> a [^] (5::nat) \<otimes> b [^] (4::nat) \<oplus>
+    a [^] (4::nat) \<otimes> b [^] (5::nat) \<oplus> a [^] (3::nat) \<otimes> b [^] (6::nat) \<oplus>
+    a [^] (2::nat) \<otimes> b [^] (7::nat) \<oplus> a \<otimes> b [^] (8::nat) \<oplus> b [^] (9::nat))"
   by ring
 
 lemma "(x::'a::field) \<noteq> 0 \<Longrightarrow> (1 - 1 / x) * x - x + 1 = 0"
--- a/src/HOL/Number_Theory/Prime_Powers.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Number_Theory/Prime_Powers.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -510,4 +510,4 @@
   finally show ?thesis .
 qed (insert assms, auto simp: mangoldt_def)
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Number_Theory/Residues.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Number_Theory/Residues.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -149,7 +149,7 @@
   using m_gt_one by (auto simp: R_def residue_ring_def)
 
 (* FIXME revise algebra library to use 1? *)
-lemma pow_cong: "(x mod m) (^) n = x^n mod m"
+lemma pow_cong: "(x mod m) [^] n = x^n mod m"
   using m_gt_one
   apply (induct n)
   apply (auto simp add: nat_pow_def one_cong)
@@ -413,12 +413,12 @@
   have car: "carrier (residue_ring (int p)) - {\<zero>\<^bsub>residue_ring (int p)\<^esub>} = {1 .. int p - 1}"
     by (auto simp add: R.zero_cong R.res_carrier_eq)
 
-  have "x (^)\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)"
+  have "x [^]\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)"
     if "x \<in> {1 .. int p - 1}" for x and i :: nat
     using that R.pow_cong[of x i] by auto
   moreover
   obtain a where a: "a \<in> {1 .. int p - 1}"
-    and a_gen: "{1 .. int p - 1} = {a(^)\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \<in> UNIV}"
+    and a_gen: "{1 .. int p - 1} = {a[^]\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \<in> UNIV}"
     using field.finite_field_mult_group_has_gen[OF R.is_field]
     by (auto simp add: car[symmetric] carrier_mult_of)
   moreover
--- a/src/HOL/Quotient_Examples/Int_Pow.thy	Fri Jan 05 15:24:57 2018 +0100
+++ b/src/HOL/Quotient_Examples/Int_Pow.thy	Fri Jan 05 18:41:42 2018 +0100
@@ -20,7 +20,7 @@
 (* first some additional lemmas that are missing in monoid *)
 
 lemma Units_nat_pow_Units [intro, simp]:
-  "a \<in> Units G \<Longrightarrow> a (^) (c :: nat) \<in> Units G" by (induct c) auto
+  "a \<in> Units G \<Longrightarrow> a [^] (c :: nat) \<in> Units G" by (induct c) auto
 
 lemma Units_r_cancel [simp]:
   "[| z \<in> Units G; x \<in> carrier G; y \<in> carrier G |] ==>
@@ -49,14 +49,14 @@
 
 lemma mult_same_comm:
   assumes [simp, intro]: "a \<in> Units G"
-  shows "a (^) (m::nat) \<otimes> inv (a (^) (n::nat)) = inv (a (^) n) \<otimes> a (^) m"
+  shows "a [^] (m::nat) \<otimes> inv (a [^] (n::nat)) = inv (a [^] n) \<otimes> a [^] m"
 proof (cases "m\<ge>n")
   have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
   case True
     then obtain k where *:"m = k + n" and **:"m = n + k" by (metis le_iff_add add.commute)
-    have "a (^) (m::nat) \<otimes> inv (a (^) (n::nat)) = a (^) k"
+    have "a [^] (m::nat) \<otimes> inv (a [^] (n::nat)) = a [^] k"
       using * by (auto simp add: nat_pow_mult[symmetric] m_assoc)
-    also have "\<dots> = inv (a (^) n) \<otimes> a (^) m"
+    also have "\<dots> = inv (a [^] n) \<otimes> a [^] m"
       using ** by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric])
     finally show ?thesis .
 next
@@ -64,15 +64,15 @@
   case False
     then obtain k where *:"n = k + m" and **:"n = m + k"
       by (metis le_iff_add add.commute nat_le_linear)
-    have "a (^) (m::nat) \<otimes> inv (a (^) (n::nat)) = inv(a (^) k)"
+    have "a [^] (m::nat) \<otimes> inv (a [^] (n::nat)) = inv(a [^] k)"
       using * by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units)
-    also have "\<dots> = inv (a (^) n) \<otimes> a (^) m"
+    also have "\<dots> = inv (a [^] n) \<otimes> a [^] m"
       using ** by (auto simp add: nat_pow_mult[symmetric] m_assoc inv_mult_units)
     finally show ?thesis .
 qed
 
 lemma mult_inv_same_comm:
-  "a \<in> Units G \<Longrightarrow> inv (a (^) (m::nat)) \<otimes> inv (a (^) (n::nat)) = inv (a (^) n) \<otimes> inv (a (^) m)"
+  "a \<in> Units G \<Longrightarrow> inv (a [^] (m::nat)) \<otimes> inv (a [^] (n::nat)) = inv (a [^] n) \<otimes> inv (a [^] m)"
 by (simp add: inv_mult_units[symmetric] nat_pow_mult ac_simps Units_closed)
 
 context
@@ -82,15 +82,15 @@
 lemma int_pow_rsp:
   assumes eq: "(b::nat) + e = d + c"
   assumes a_in_G [simp, intro]: "a \<in> Units G"
-  shows "a (^) b \<otimes> inv (a (^) c) = a (^) d \<otimes> inv (a (^) e)"
+  shows "a [^] b \<otimes> inv (a [^] c) = a [^] d \<otimes> inv (a [^] e)"
 proof(cases "b\<ge>c")
   have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
   case True
     then obtain n where "b = n + c" by (metis le_iff_add add.commute)
     then have "d = n + e" using eq by arith
-    from \<open>b = _\<close> have "a (^) b \<otimes> inv (a (^) c) = a (^) n"
+    from \<open>b = _\<close> have "a [^] b \<otimes> inv (a [^] c) = a [^] n"
       by (auto simp add: nat_pow_mult[symmetric] m_assoc)
-    also from \<open>d = _\<close>  have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
+    also from \<open>d = _\<close>  have "\<dots> = a [^] d \<otimes> inv (a [^] e)"
       by (auto simp add: nat_pow_mult[symmetric] m_assoc)
     finally show ?thesis .
 next
@@ -98,20 +98,20 @@
   case False
     then obtain n where "c = n + b" by (metis le_iff_add add.commute nat_le_linear)
     then have "e = n + d" using eq by arith
-    from \<open>c = _\<close> have "a (^) b \<otimes> inv (a (^) c) = inv (a (^) n)"
+    from \<open>c = _\<close> have "a [^] b \<otimes> inv (a [^] c) = inv (a [^] n)"
       by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units)
-    also from \<open>e = _\<close> have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
+    also from \<open>e = _\<close> have "\<dots> = a [^] d \<otimes> inv (a [^] e)"
       by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units)
     finally show ?thesis .
 qed
 
 (*
   This definition is more convinient than the definition in HOL/Algebra/Group because
-  it doesn't contain a test z < 0 when a (^) z is being defined.
+  it doesn't contain a test z < 0 when a [^] z is being defined.
 *)
 
 lift_definition int_pow :: "('a, 'm) monoid_scheme \<Rightarrow> 'a \<Rightarrow> int \<Rightarrow> 'a" is
-  "\<lambda>G a (n1, n2). if a \<in> Units G \<and> monoid G then (a (^)\<^bsub>G\<^esub> n1) \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> (a (^)\<^bsub>G\<^esub> n2)) else \<one>\<^bsub>G\<^esub>"
+  "\<lambda>G a (n1, n2). if a \<in> Units G \<and> monoid G then (a [^]\<^bsub>G\<^esub> n1) \<otimes>\<^bsub>G\<^esub> (inv\<^bsub>G\<^esub> (a [^]\<^bsub>G\<^esub> n2)) else \<one>\<^bsub>G\<^esub>"
 unfolding intrel_def by (auto intro: monoid.int_pow_rsp)
 
 (*
@@ -125,12 +125,12 @@
 proof -
   {
     fix k l m :: nat
-    have "a (^) l \<otimes> (inv (a (^) m) \<otimes> inv (a (^) k)) = (a (^) l \<otimes> inv (a (^) k)) \<otimes> inv (a (^) m)"
+    have "a [^] l \<otimes> (inv (a [^] m) \<otimes> inv (a [^] k)) = (a [^] l \<otimes> inv (a [^] k)) \<otimes> inv (a [^] m)"
       (is "?lhs = _")
       by (simp add: mult_inv_same_comm m_assoc Units_closed)
-    also have "\<dots> = (inv (a (^) k) \<otimes> a (^) l) \<otimes> inv (a (^) m)"
+    also have "\<dots> = (inv (a [^] k) \<otimes> a [^] l) \<otimes> inv (a [^] m)"
       by (simp add: mult_same_comm)
-    also have "\<dots> = inv (a (^) k) \<otimes> (a (^) l \<otimes> inv (a (^) m))" (is "_ = ?rhs")
+    also have "\<dots> = inv (a [^] k) \<otimes> (a [^] l \<otimes> inv (a [^] m))" (is "_ = ?rhs")
       by (simp add: m_assoc Units_closed)
     finally have "?lhs = ?rhs" .
   }