src/HOL/Library/Lexord.thy
author wenzelm
Fri, 20 Sep 2024 19:51:08 +0200
changeset 80914 d97fdabd9e2b
parent 77955 c4677a6aae2c
child 82774 2865a6618cba
permissions -rw-r--r--
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73794
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     1
section \<open>Lexicographic orderings\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     2
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     3
theory Lexord
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     4
  imports Main
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     5
begin
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     6
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     7
subsection \<open>The preorder case\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     8
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
     9
locale lex_preordering = preordering
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    10
begin
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    11
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    12
inductive lex_less :: \<open>'a list \<Rightarrow> 'a list \<Rightarrow> bool\<close>  (infix \<open>[\<^bold><]\<close> 50) 
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    13
where
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    14
  Nil: \<open>[] [\<^bold><] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    15
| Cons: \<open>x \<^bold>< y \<Longrightarrow> x # xs [\<^bold><] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    16
| Cons_eq: \<open>x \<^bold>\<le> y \<Longrightarrow> y \<^bold>\<le> x \<Longrightarrow> xs [\<^bold><] ys \<Longrightarrow> x # xs [\<^bold><] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    17
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    18
inductive lex_less_eq :: \<open>'a list \<Rightarrow> 'a list \<Rightarrow> bool\<close>  (infix \<open>[\<^bold>\<le>]\<close> 50)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    19
where
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    20
  Nil: \<open>[] [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    21
| Cons: \<open>x \<^bold>< y \<Longrightarrow> x # xs [\<^bold>\<le>] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    22
| Cons_eq: \<open>x \<^bold>\<le> y \<Longrightarrow> y \<^bold>\<le> x \<Longrightarrow> xs [\<^bold>\<le>] ys \<Longrightarrow> x # xs [\<^bold>\<le>] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    23
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    24
lemma lex_less_simps [simp]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    25
  \<open>[] [\<^bold><] y # ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    26
  \<open>\<not> xs [\<^bold><] []\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    27
  \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold><] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    28
  by (auto intro: lex_less.intros elim: lex_less.cases)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    29
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    30
lemma lex_less_eq_simps [simp]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    31
  \<open>[] [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    32
  \<open>\<not> x # xs [\<^bold>\<le>] []\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    33
  \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    34
  by (auto intro: lex_less_eq.intros elim: lex_less_eq.cases)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    35
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    36
lemma lex_less_code [code]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    37
  \<open>[] [\<^bold><] y # ys \<longleftrightarrow> True\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    38
  \<open>xs [\<^bold><] [] \<longleftrightarrow> False\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    39
  \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold><] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    40
  by simp_all
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    41
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    42
lemma lex_less_eq_code [code]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    43
  \<open>[] [\<^bold>\<le>] ys \<longleftrightarrow> True\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    44
  \<open>x # xs [\<^bold>\<le>] [] \<longleftrightarrow> False\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    45
  \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x \<^bold>\<le> y \<and> y \<^bold>\<le> x \<and> xs [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    46
  by simp_all
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    47
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    48
lemma preordering:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    49
  \<open>preordering ([\<^bold>\<le>]) ([\<^bold><])\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    50
proof
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    51
  fix xs ys zs
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    52
  show \<open>xs [\<^bold>\<le>] xs\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    53
    by (induction xs) (simp_all add: refl)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    54
  show \<open>xs [\<^bold>\<le>] zs\<close> if \<open>xs [\<^bold>\<le>] ys\<close> \<open>ys [\<^bold>\<le>] zs\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    55
  using that proof (induction arbitrary: zs)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    56
    case (Nil ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    57
    then show ?case by simp
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    58
  next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    59
    case (Cons x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    60
    then show ?case
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    61
      by (cases zs) (auto dest: strict_trans strict_trans2)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    62
  next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    63
    case (Cons_eq x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    64
    then show ?case
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    65
      by (cases zs) (auto dest: strict_trans1 intro: trans)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    66
  qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    67
  show \<open>xs [\<^bold><] ys \<longleftrightarrow> xs [\<^bold>\<le>] ys \<and> \<not> ys [\<^bold>\<le>] xs\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    68
  proof
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    69
    assume ?P
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    70
    then have \<open>xs [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    71
      by induction simp_all
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    72
    moreover have \<open>\<not> ys [\<^bold>\<le>] xs\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    73
      using \<open>?P\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    74
      by induction (simp_all, simp_all add: strict_iff_not asym)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    75
    ultimately show ?Q ..
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    76
  next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    77
    assume ?Q
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    78
    then have \<open>xs [\<^bold>\<le>] ys\<close> \<open>\<not> ys [\<^bold>\<le>] xs\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    79
      by auto
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    80
    then show ?P
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    81
    proof induction
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    82
      case (Nil ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    83
      then show ?case
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    84
        by (cases ys) simp_all
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    85
    next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    86
      case (Cons x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    87
      then show ?case
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    88
        by simp
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    89
    next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    90
      case (Cons_eq x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    91
      then show ?case
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    92
        by simp
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    93
    qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    94
  qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    95
qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    96
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    97
interpretation lex: preordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    98
  by (fact preordering)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
    99
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   100
end
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   101
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   102
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   103
subsection \<open>The order case\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   104
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   105
locale lex_ordering = lex_preordering + ordering
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   106
begin
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   107
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   108
interpretation lex: preordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   109
  by (fact preordering)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   110
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   111
lemma less_lex_Cons_iff [simp]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   112
  \<open>x # xs [\<^bold><] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x = y \<and> xs [\<^bold><] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   113
  by (auto intro: refl antisym)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   114
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   115
lemma less_eq_lex_Cons_iff [simp]:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   116
  \<open>x # xs [\<^bold>\<le>] y # ys \<longleftrightarrow> x \<^bold>< y \<or> x = y \<and> xs [\<^bold>\<le>] ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   117
  by (auto intro: refl antisym)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   118
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   119
lemma ordering:
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   120
  \<open>ordering ([\<^bold>\<le>]) ([\<^bold><])\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   121
proof
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   122
  fix xs ys
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   123
  show *: \<open>xs = ys\<close> if \<open>xs [\<^bold>\<le>] ys\<close> \<open>ys [\<^bold>\<le>] xs\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   124
  using that proof induction
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   125
  case (Nil ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   126
    then show ?case by (cases ys) simp
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   127
  next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   128
    case (Cons x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   129
    then show ?case by (auto dest: asym intro: antisym)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   130
      (simp add: strict_iff_not)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   131
  next
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   132
    case (Cons_eq x y xs ys)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   133
    then show ?case by (auto intro: antisym)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   134
      (simp add: strict_iff_not)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   135
  qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   136
  show \<open>xs [\<^bold><] ys \<longleftrightarrow> xs [\<^bold>\<le>] ys \<and> xs \<noteq> ys\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   137
    by (auto simp add: lex.strict_iff_not dest: *)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   138
qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   139
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   140
interpretation lex: ordering \<open>([\<^bold>\<le>])\<close> \<open>([\<^bold><])\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   141
  by (fact ordering)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   142
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   143
end
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   144
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   145
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   146
subsection \<open>Canonical instance\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   147
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   148
instantiation list :: (preorder) preorder
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   149
begin
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   150
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   151
global_interpretation lex: lex_preordering \<open>(\<le>) :: 'a::preorder \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   152
  defines less_eq_list = lex.lex_less_eq
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   153
    and less_list = lex.lex_less ..
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   154
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   155
instance
77955
c4677a6aae2c more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
wenzelm
parents: 75244
diff changeset
   156
  by (rule preorder.intro_of_class, rule preordering_preorderI, fact lex.preordering)
73794
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   157
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   158
end
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   159
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   160
global_interpretation lex: lex_ordering \<open>(\<le>) :: 'a::order \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   161
  rewrites \<open>lex_preordering.lex_less_eq (\<le>) (<) = ((\<le>) :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool)\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   162
    and \<open>lex_preordering.lex_less (\<le>) (<) = ((<) :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool)\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   163
proof -
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   164
  interpret lex_ordering \<open>(\<le>) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> \<open>(<) :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> ..
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   165
  show \<open>lex_ordering ((\<le>)  :: 'a \<Rightarrow> 'a \<Rightarrow> bool) (<)\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   166
    by (fact lex_ordering_axioms)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   167
  show \<open>lex_preordering.lex_less_eq (\<le>) (<) = (\<le>)\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   168
    by (simp add: less_eq_list_def)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   169
  show \<open>lex_preordering.lex_less (\<le>) (<) = (<)\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   170
    by (simp add: less_list_def)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   171
qed
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   172
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   173
instance list :: (order) order
77955
c4677a6aae2c more standard name bindings (amending 5bf71b4da706): avoid odd full_name like "Orderings.class.Orderings.preorder.of_class.intro" with many redundant name space accesses;
wenzelm
parents: 75244
diff changeset
   174
  by (rule order.intro_of_class, rule ordering_orderI, fact lex.ordering)
73794
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   175
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   176
export_code \<open>(\<le>) :: _ list \<Rightarrow> _ list \<Rightarrow> bool\<close> \<open>(<) :: _ list \<Rightarrow> _ list \<Rightarrow> bool\<close> in Haskell
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   177
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   178
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   179
subsection \<open>Non-canonical instance\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   180
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   181
context comm_monoid_mult
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   182
begin
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   183
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   184
definition dvd_strict :: \<open>'a \<Rightarrow> 'a \<Rightarrow> bool\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   185
  where \<open>dvd_strict a b \<longleftrightarrow> a dvd b \<and> \<not> b dvd a\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   186
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   187
end
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   188
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   189
global_interpretation dvd: lex_preordering \<open>(dvd) :: 'a::comm_monoid_mult \<Rightarrow> 'a \<Rightarrow> bool\<close> dvd_strict
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   190
  defines lex_dvd = dvd.lex_less_eq
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   191
    and lex_dvd_strict = dvd.lex_less
75244
f70b1a2c2783 A tiny further cleanup
paulson <lp15@cam.ac.uk>
parents: 73846
diff changeset
   192
  by unfold_locales (auto simp add: dvd_strict_def)
73794
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   193
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   194
global_interpretation lex_dvd: preordering lex_dvd lex_dvd_strict
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   195
  by (fact dvd.preordering)
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   196
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   197
definition \<open>example = lex_dvd [(4::int), - 7, 8] [- 8, 13, 5]\<close>
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   198
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   199
export_code example in Haskell
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   200
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   201
value example
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   202
e75635a0bafd lexorders the locale way
haftmann
parents:
diff changeset
   203
end