src/HOL/Examples/Adhoc_Overloading_Examples.thy
changeset 80914 d97fdabd9e2b
parent 72029 83456d9f0ed5
equal deleted inserted replaced
80913:46f59511b7bb 80914:d97fdabd9e2b
   174 
   174 
   175 text \<open>We want to be able to apply permutations to arbitrary types. To
   175 text \<open>We want to be able to apply permutations to arbitrary types. To
   176 this end we introduce a constant \<open>PERMUTE\<close> together with
   176 this end we introduce a constant \<open>PERMUTE\<close> together with
   177 convenient infix syntax.\<close>
   177 convenient infix syntax.\<close>
   178 
   178 
   179 consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr "\<bullet>" 75)
   179 consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr \<open>\<bullet>\<close> 75)
   180 
   180 
   181 text \<open>Then we add a locale for types \<^typ>\<open>'b\<close> that support
   181 text \<open>Then we add a locale for types \<^typ>\<open>'b\<close> that support
   182 appliciation of permutations.\<close>
   182 appliciation of permutations.\<close>
   183 locale permute =
   183 locale permute =
   184   fixes permute :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"
   184   fixes permute :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"