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