src/HOL/Examples/Adhoc_Overloading.thy
changeset 81989 96afb0707532
parent 80914 d97fdabd9e2b
child 81995 d67dadd69d07
equal deleted inserted replaced
81988:846293abd12d 81989:96afb0707532
       
     1 (*  Title:      HOL/Examples/Adhoc_Overloading.thy
       
     2     Author:     Christian Sternagel
       
     3 *)
       
     4 
       
     5 section \<open>Ad Hoc Overloading\<close>
       
     6 
       
     7 theory Adhoc_Overloading
       
     8 imports
       
     9   Main
       
    10   "HOL-Library.Infinite_Set"
       
    11 begin
       
    12 
       
    13 text \<open>Adhoc overloading allows to overload a constant depending on
       
    14 its type. Typically this involves to introduce an uninterpreted
       
    15 constant (used for input and output) and then add some variants (used
       
    16 internally).\<close>
       
    17 
       
    18 subsection \<open>Plain Ad Hoc Overloading\<close>
       
    19 
       
    20 text \<open>Consider the type of first-order terms.\<close>
       
    21 datatype ('a, 'b) "term" =
       
    22   Var 'b |
       
    23   Fun 'a "('a, 'b) term list"
       
    24 
       
    25 text \<open>The set of variables of a term might be computed as follows.\<close>
       
    26 fun term_vars :: "('a, 'b) term \<Rightarrow> 'b set" where
       
    27   "term_vars (Var x) = {x}" |
       
    28   "term_vars (Fun f ts) = \<Union>(set (map term_vars ts))"
       
    29 
       
    30 text \<open>However, also for \emph{rules} (i.e., pairs of terms) and term
       
    31 rewrite systems (i.e., sets of rules), the set of variables makes
       
    32 sense. Thus we introduce an unspecified constant \<open>vars\<close>.\<close>
       
    33 
       
    34 consts vars :: "'a \<Rightarrow> 'b set"
       
    35 
       
    36 text \<open>Which is then overloaded with variants for terms, rules, and TRSs.\<close>
       
    37 adhoc_overloading
       
    38   vars term_vars
       
    39 
       
    40 value [nbe] "vars (Fun ''f'' [Var 0, Var 1])"
       
    41 
       
    42 fun rule_vars :: "('a, 'b) term \<times> ('a, 'b) term \<Rightarrow> 'b set" where
       
    43   "rule_vars (l, r) = vars l \<union> vars r"
       
    44 
       
    45 adhoc_overloading
       
    46   vars rule_vars
       
    47 
       
    48 value [nbe] "vars (Var 1, Var 0)"
       
    49 
       
    50 definition trs_vars :: "(('a, 'b) term \<times> ('a, 'b) term) set \<Rightarrow> 'b set" where
       
    51   "trs_vars R = \<Union>(rule_vars ` R)"
       
    52 
       
    53 adhoc_overloading
       
    54   vars trs_vars
       
    55 
       
    56 value [nbe] "vars {(Var 1, Var 0)}"
       
    57 
       
    58 text \<open>Sometimes it is necessary to add explicit type constraints
       
    59 before a variant can be determined.\<close>
       
    60 (*value "vars R" (*has multiple instances*)*)
       
    61 value "vars (R :: (('a, 'b) term \<times> ('a, 'b) term) set)"
       
    62 
       
    63 text \<open>It is also possible to remove variants.\<close>
       
    64 no_adhoc_overloading
       
    65   vars term_vars rule_vars 
       
    66 
       
    67 (*value "vars (Var 1)" (*does not have an instance*)*)
       
    68 
       
    69 text \<open>As stated earlier, the overloaded constant is only used for
       
    70 input and output. Internally, always a variant is used, as can be
       
    71 observed by the configuration option \<open>show_variants\<close>.\<close>
       
    72 
       
    73 adhoc_overloading
       
    74   vars term_vars
       
    75 
       
    76 declare [[show_variants]]
       
    77 
       
    78 term "vars (Var 1)" (*which yields: "term_vars (Var 1)"*)
       
    79 
       
    80 
       
    81 subsection \<open>Adhoc Overloading inside Locales\<close>
       
    82 
       
    83 text \<open>As example we use permutations that are parametrized over an
       
    84 atom type \<^typ>\<open>'a\<close>.\<close>
       
    85 
       
    86 definition perms :: "('a \<Rightarrow> 'a) set" where
       
    87   "perms = {f. bij f \<and> finite {x. f x \<noteq> x}}"
       
    88 
       
    89 typedef 'a perm = "perms :: ('a \<Rightarrow> 'a) set"
       
    90   by standard (auto simp: perms_def)
       
    91 
       
    92 text \<open>First we need some auxiliary lemmas.\<close>
       
    93 lemma permsI [Pure.intro]:
       
    94   assumes "bij f" and "MOST x. f x = x"
       
    95   shows "f \<in> perms"
       
    96   using assms by (auto simp: perms_def) (metis MOST_iff_finiteNeg)
       
    97 
       
    98 lemma perms_imp_bij:
       
    99   "f \<in> perms \<Longrightarrow> bij f"
       
   100   by (simp add: perms_def)
       
   101 
       
   102 lemma perms_imp_MOST_eq:
       
   103   "f \<in> perms \<Longrightarrow> MOST x. f x = x"
       
   104   by (simp add: perms_def) (metis MOST_iff_finiteNeg)
       
   105 
       
   106 lemma id_perms [simp]:
       
   107   "id \<in> perms"
       
   108   "(\<lambda>x. x) \<in> perms"
       
   109   by (auto simp: perms_def bij_def)
       
   110 
       
   111 lemma perms_comp [simp]:
       
   112   assumes f: "f \<in> perms" and g: "g \<in> perms"
       
   113   shows "(f \<circ> g) \<in> perms"
       
   114   apply (intro permsI bij_comp)
       
   115   apply (rule perms_imp_bij [OF g])
       
   116   apply (rule perms_imp_bij [OF f])
       
   117   apply (rule MOST_rev_mp [OF perms_imp_MOST_eq [OF g]])
       
   118   apply (rule MOST_rev_mp [OF perms_imp_MOST_eq [OF f]])
       
   119   by simp
       
   120 
       
   121 lemma perms_inv:
       
   122   assumes f: "f \<in> perms"
       
   123   shows "inv f \<in> perms"
       
   124   apply (rule permsI)
       
   125   apply (rule bij_imp_bij_inv)
       
   126   apply (rule perms_imp_bij [OF f])
       
   127   apply (rule MOST_mono [OF perms_imp_MOST_eq [OF f]])
       
   128   apply (erule subst, rule inv_f_f)
       
   129   apply (rule bij_is_inj [OF perms_imp_bij [OF f]])
       
   130   done
       
   131 
       
   132 lemma bij_Rep_perm: "bij (Rep_perm p)"
       
   133   using Rep_perm [of p] unfolding perms_def by simp
       
   134 
       
   135 instantiation perm :: (type) group_add
       
   136 begin
       
   137 
       
   138 definition "0 = Abs_perm id"
       
   139 definition "- p = Abs_perm (inv (Rep_perm p))"
       
   140 definition "p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)"
       
   141 definition "(p1::'a perm) - p2 = p1 + - p2"
       
   142 
       
   143 lemma Rep_perm_0: "Rep_perm 0 = id"
       
   144   unfolding zero_perm_def by (simp add: Abs_perm_inverse)
       
   145 
       
   146 lemma Rep_perm_add:
       
   147   "Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2"
       
   148   unfolding plus_perm_def by (simp add: Abs_perm_inverse Rep_perm)
       
   149 
       
   150 lemma Rep_perm_uminus:
       
   151   "Rep_perm (- p) = inv (Rep_perm p)"
       
   152   unfolding uminus_perm_def by (simp add: Abs_perm_inverse perms_inv Rep_perm)
       
   153 
       
   154 instance
       
   155   apply standard
       
   156   unfolding Rep_perm_inject [symmetric]
       
   157   unfolding minus_perm_def
       
   158   unfolding Rep_perm_add
       
   159   unfolding Rep_perm_uminus
       
   160   unfolding Rep_perm_0
       
   161   apply (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]])
       
   162   done
       
   163 
       
   164 end
       
   165 
       
   166 lemmas Rep_perm_simps =
       
   167   Rep_perm_0
       
   168   Rep_perm_add
       
   169   Rep_perm_uminus
       
   170 
       
   171 
       
   172 section \<open>Permutation Types\<close>
       
   173 
       
   174 text \<open>We want to be able to apply permutations to arbitrary types. To
       
   175 this end we introduce a constant \<open>PERMUTE\<close> together with
       
   176 convenient infix syntax.\<close>
       
   177 
       
   178 consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr \<open>\<bullet>\<close> 75)
       
   179 
       
   180 text \<open>Then we add a locale for types \<^typ>\<open>'b\<close> that support
       
   181 appliciation of permutations.\<close>
       
   182 locale permute =
       
   183   fixes permute :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"
       
   184   assumes permute_zero [simp]: "permute 0 x = x"
       
   185     and permute_plus [simp]: "permute (p + q) x = permute p (permute q x)"
       
   186 begin
       
   187 
       
   188 adhoc_overloading
       
   189   PERMUTE permute
       
   190 
       
   191 end
       
   192 
       
   193 text \<open>Permuting atoms.\<close>
       
   194 definition permute_atom :: "'a perm \<Rightarrow> 'a \<Rightarrow> 'a" where
       
   195   "permute_atom p a = (Rep_perm p) a"
       
   196 
       
   197 adhoc_overloading
       
   198   PERMUTE permute_atom
       
   199 
       
   200 interpretation atom_permute: permute permute_atom
       
   201   by standard (simp_all add: permute_atom_def Rep_perm_simps)
       
   202 
       
   203 text \<open>Permuting permutations.\<close>
       
   204 definition permute_perm :: "'a perm \<Rightarrow> 'a perm \<Rightarrow> 'a perm" where
       
   205   "permute_perm p q = p + q - p"
       
   206 
       
   207 adhoc_overloading
       
   208   PERMUTE permute_perm
       
   209 
       
   210 interpretation perm_permute: permute permute_perm
       
   211   apply standard
       
   212   unfolding permute_perm_def
       
   213   apply simp
       
   214   apply (simp only: diff_conv_add_uminus minus_add add.assoc)
       
   215   done
       
   216 
       
   217 text \<open>Permuting functions.\<close>
       
   218 locale fun_permute =
       
   219   dom: permute perm1 + ran: permute perm2
       
   220   for perm1 :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b"
       
   221   and perm2 :: "'a perm \<Rightarrow> 'c \<Rightarrow> 'c"
       
   222 begin
       
   223 
       
   224 adhoc_overloading
       
   225   PERMUTE perm1 perm2
       
   226 
       
   227 definition permute_fun :: "'a perm \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c)" where
       
   228   "permute_fun p f = (\<lambda>x. p \<bullet> (f (-p \<bullet> x)))"
       
   229 
       
   230 adhoc_overloading
       
   231   PERMUTE permute_fun
       
   232 
       
   233 end
       
   234 
       
   235 sublocale fun_permute \<subseteq> permute permute_fun
       
   236   by (unfold_locales, auto simp: permute_fun_def)
       
   237      (metis dom.permute_plus minus_add)
       
   238 
       
   239 lemma "(Abs_perm id :: nat perm) \<bullet> Suc 0 = Suc 0"
       
   240   unfolding permute_atom_def
       
   241   by (metis Rep_perm_0 id_apply zero_perm_def)
       
   242 
       
   243 interpretation atom_fun_permute: fun_permute permute_atom permute_atom
       
   244   by (unfold_locales)
       
   245 
       
   246 adhoc_overloading
       
   247   PERMUTE atom_fun_permute.permute_fun
       
   248 
       
   249 lemma "(Abs_perm id :: 'a perm) \<bullet> id = id"
       
   250   unfolding atom_fun_permute.permute_fun_def
       
   251   unfolding permute_atom_def
       
   252   by (metis Rep_perm_0 id_def inj_imp_inv_eq inj_on_id uminus_perm_def zero_perm_def)
       
   253 
       
   254 end