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