| 
27701
 | 
     1  | 
(*
  | 
| 
 | 
     2  | 
  Title:     Divisibility in monoids and rings
  | 
| 
 | 
     3  | 
  Id:        $Id$
  | 
| 
 | 
     4  | 
  Author:    Clemens Ballarin, started 18 July 2008
  | 
| 
 | 
     5  | 
  Copyright: Clemens Ballarin
  | 
| 
 | 
     6  | 
*)
  | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory Divisibility
  | 
| 
 | 
     9  | 
imports Permutation Coset Group GLattice
  | 
| 
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
subsection {* Monoid with cancelation law *}
 | 
| 
 | 
    13  | 
  | 
| 
 | 
    14  | 
locale monoid_cancel = monoid +
  | 
| 
 | 
    15  | 
  assumes l_cancel: 
  | 
| 
 | 
    16  | 
          "\<lbrakk>c \<otimes> a = c \<otimes> b; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
  | 
| 
 | 
    17  | 
      and r_cancel: 
  | 
| 
 | 
    18  | 
          "\<lbrakk>a \<otimes> c = b \<otimes> c; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
  | 
| 
 | 
    19  | 
  | 
| 
 | 
    20  | 
lemma (in monoid) monoid_cancelI:
  | 
| 
 | 
    21  | 
  assumes l_cancel: 
  | 
| 
 | 
    22  | 
          "\<And>a b c. \<lbrakk>c \<otimes> a = c \<otimes> b; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
  | 
| 
 | 
    23  | 
      and r_cancel: 
  | 
| 
 | 
    24  | 
          "\<And>a b c. \<lbrakk>a \<otimes> c = b \<otimes> c; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
  | 
| 
 | 
    25  | 
  shows "monoid_cancel G"
  | 
| 
 | 
    26  | 
by unfold_locales fact+
  | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
lemma (in monoid_cancel) is_monoid_cancel:
  | 
| 
 | 
    29  | 
  "monoid_cancel G"
  | 
| 
 | 
    30  | 
by intro_locales
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
interpretation group \<subseteq> monoid_cancel
  | 
| 
 | 
    33  | 
by unfold_locales simp+
  | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
locale comm_monoid_cancel = monoid_cancel + comm_monoid
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
lemma comm_monoid_cancelI:
  | 
| 
 | 
    39  | 
  includes comm_monoid
  | 
| 
 | 
    40  | 
  assumes cancel: 
  | 
| 
 | 
    41  | 
          "\<And>a b c. \<lbrakk>a \<otimes> c = b \<otimes> c; a \<in> carrier G; b \<in> carrier G; c \<in> carrier G\<rbrakk> \<Longrightarrow> a = b"
  | 
| 
 | 
    42  | 
  shows "comm_monoid_cancel G"
  | 
| 
 | 
    43  | 
apply unfold_locales
  | 
| 
 | 
    44  | 
 apply (subgoal_tac "a \<otimes> c = b \<otimes> c")
  | 
| 
 | 
    45  | 
  apply (iprover intro: cancel)
  | 
| 
 | 
    46  | 
 apply (simp add: m_comm)
  | 
| 
 | 
    47  | 
apply (iprover intro: cancel)
  | 
| 
 | 
    48  | 
done
  | 
| 
 | 
    49  | 
  | 
| 
 | 
    50  | 
lemma (in comm_monoid_cancel) is_comm_monoid_cancel:
  | 
| 
 | 
    51  | 
  "comm_monoid_cancel G"
  | 
| 
 | 
    52  | 
by intro_locales
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
interpretation comm_group \<subseteq> comm_monoid_cancel
  | 
| 
 | 
    55  | 
by unfold_locales
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
subsection {* Products of units in monoids *}
 | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
lemma (in monoid) Units_m_closed[simp, intro]:
  | 
| 
 | 
    61  | 
  assumes h1unit: "h1 \<in> Units G" and h2unit: "h2 \<in> Units G"
  | 
| 
 | 
    62  | 
  shows "h1 \<otimes> h2 \<in> Units G"
  | 
| 
 | 
    63  | 
unfolding Units_def
  | 
| 
 | 
    64  | 
using assms
  | 
| 
 | 
    65  | 
apply safe
  | 
| 
 | 
    66  | 
apply fast
  | 
| 
 | 
    67  | 
apply (intro bexI[of _ "inv h2 \<otimes> inv h1"], safe)
  | 
| 
 | 
    68  | 
  apply (simp add: m_assoc Units_closed)
  | 
| 
 | 
    69  | 
  apply (simp add: m_assoc[symmetric] Units_closed Units_l_inv)
  | 
| 
 | 
    70  | 
 apply (simp add: m_assoc Units_closed)
  | 
| 
 | 
    71  | 
 apply (simp add: m_assoc[symmetric] Units_closed Units_r_inv)
  | 
| 
 | 
    72  | 
apply fast
  | 
| 
 | 
    73  | 
done
  | 
| 
 | 
    74  | 
  | 
| 
 | 
    75  | 
lemma (in monoid) prod_unit_l:
  | 
| 
 | 
    76  | 
  assumes abunit[simp]: "a \<otimes> b \<in> Units G" and aunit[simp]: "a \<in> Units G"
  | 
| 
 | 
    77  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
    78  | 
  shows "b \<in> Units G"
  | 
| 
 | 
    79  | 
proof -
  | 
| 
 | 
    80  | 
  have c: "inv (a \<otimes> b) \<otimes> a \<in> carrier G" by simp
  | 
| 
 | 
    81  | 
  | 
| 
 | 
    82  | 
  have "(inv (a \<otimes> b) \<otimes> a) \<otimes> b = inv (a \<otimes> b) \<otimes> (a \<otimes> b)" by (simp add: m_assoc)
  | 
| 
 | 
    83  | 
  also have "\<dots> = \<one>" by (simp add: Units_l_inv)
  | 
| 
 | 
    84  | 
  finally have li: "(inv (a \<otimes> b) \<otimes> a) \<otimes> b = \<one>" .
  | 
| 
 | 
    85  | 
  | 
| 
 | 
    86  | 
  have "\<one> = inv a \<otimes> a" by (simp add: Units_l_inv[symmetric])
  | 
| 
 | 
    87  | 
  also have "\<dots> = inv a \<otimes> \<one> \<otimes> a" by simp
  | 
| 
 | 
    88  | 
  also have "\<dots> = inv a \<otimes> ((a \<otimes> b) \<otimes> inv (a \<otimes> b)) \<otimes> a"
  | 
| 
 | 
    89  | 
       by (simp add: Units_r_inv[OF abunit, symmetric] del: Units_r_inv)
  | 
| 
 | 
    90  | 
  also have "\<dots> = ((inv a \<otimes> a) \<otimes> b) \<otimes> inv (a \<otimes> b) \<otimes> a"
  | 
| 
 | 
    91  | 
    by (simp add: m_assoc del: Units_l_inv)
  | 
| 
 | 
    92  | 
  also have "\<dots> = b \<otimes> inv (a \<otimes> b) \<otimes> a" by (simp add: Units_l_inv)
  | 
| 
 | 
    93  | 
  also have "\<dots> = b \<otimes> (inv (a \<otimes> b) \<otimes> a)" by (simp add: m_assoc)
  | 
| 
 | 
    94  | 
  finally have ri: "b \<otimes> (inv (a \<otimes> b) \<otimes> a) = \<one> " by simp
  | 
| 
 | 
    95  | 
  | 
| 
 | 
    96  | 
  from c li ri
  | 
| 
 | 
    97  | 
      show "b \<in> Units G" by (simp add: Units_def, fast)
  | 
| 
 | 
    98  | 
qed
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
lemma (in monoid) prod_unit_r:
  | 
| 
 | 
   101  | 
  assumes abunit[simp]: "a \<otimes> b \<in> Units G" and bunit[simp]: "b \<in> Units G"
  | 
| 
 | 
   102  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   103  | 
  shows "a \<in> Units G"
  | 
| 
 | 
   104  | 
proof -
  | 
| 
 | 
   105  | 
  have c: "b \<otimes> inv (a \<otimes> b) \<in> carrier G" by simp
  | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
  have "a \<otimes> (b \<otimes> inv (a \<otimes> b)) = (a \<otimes> b) \<otimes> inv (a \<otimes> b)"
  | 
| 
 | 
   108  | 
    by (simp add: m_assoc del: Units_r_inv)
  | 
| 
 | 
   109  | 
  also have "\<dots> = \<one>" by simp
  | 
| 
 | 
   110  | 
  finally have li: "a \<otimes> (b \<otimes> inv (a \<otimes> b)) = \<one>" .
  | 
| 
 | 
   111  | 
  | 
| 
 | 
   112  | 
  have "\<one> = b \<otimes> inv b" by (simp add: Units_r_inv[symmetric])
  | 
| 
 | 
   113  | 
  also have "\<dots> = b \<otimes> \<one> \<otimes> inv b" by simp
  | 
| 
 | 
   114  | 
  also have "\<dots> = b \<otimes> (inv (a \<otimes> b) \<otimes> (a \<otimes> b)) \<otimes> inv b" 
  | 
| 
 | 
   115  | 
       by (simp add: Units_l_inv[OF abunit, symmetric] del: Units_l_inv)
  | 
| 
 | 
   116  | 
  also have "\<dots> = (b \<otimes> inv (a \<otimes> b) \<otimes> a) \<otimes> (b \<otimes> inv b)"
  | 
| 
 | 
   117  | 
    by (simp add: m_assoc del: Units_l_inv)
  | 
| 
 | 
   118  | 
  also have "\<dots> = b \<otimes> inv (a \<otimes> b) \<otimes> a" by simp
  | 
| 
 | 
   119  | 
  finally have ri: "(b \<otimes> inv (a \<otimes> b)) \<otimes> a = \<one> " by simp
  | 
| 
 | 
   120  | 
  | 
| 
 | 
   121  | 
  from c li ri
  | 
| 
 | 
   122  | 
      show "a \<in> Units G" by (simp add: Units_def, fast)
  | 
| 
 | 
   123  | 
qed
  | 
| 
 | 
   124  | 
  | 
| 
 | 
   125  | 
lemma (in comm_monoid) unit_factor:
  | 
| 
 | 
   126  | 
  assumes abunit: "a \<otimes> b \<in> Units G"
  | 
| 
 | 
   127  | 
    and [simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   128  | 
  shows "a \<in> Units G"
  | 
| 
 | 
   129  | 
using abunit[simplified Units_def]
  | 
| 
 | 
   130  | 
proof clarsimp
  | 
| 
 | 
   131  | 
  fix i
  | 
| 
 | 
   132  | 
  assume [simp]: "i \<in> carrier G"
  | 
| 
 | 
   133  | 
    and li: "i \<otimes> (a \<otimes> b) = \<one>"
  | 
| 
 | 
   134  | 
    and ri: "a \<otimes> b \<otimes> i = \<one>"
  | 
| 
 | 
   135  | 
  | 
| 
 | 
   136  | 
  have carr': "b \<otimes> i \<in> carrier G" by simp
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
  have "(b \<otimes> i) \<otimes> a = (i \<otimes> b) \<otimes> a" by (simp add: m_comm)
  | 
| 
 | 
   139  | 
  also have "\<dots> = i \<otimes> (b \<otimes> a)" by (simp add: m_assoc)
  | 
| 
 | 
   140  | 
  also have "\<dots> = i \<otimes> (a \<otimes> b)" by (simp add: m_comm)
  | 
| 
 | 
   141  | 
  also note li
  | 
| 
 | 
   142  | 
  finally have li': "(b \<otimes> i) \<otimes> a = \<one>" .
  | 
| 
 | 
   143  | 
  | 
| 
 | 
   144  | 
  have "a \<otimes> (b \<otimes> i) = a \<otimes> b \<otimes> i" by (simp add: m_assoc)
  | 
| 
 | 
   145  | 
  also note ri
  | 
| 
 | 
   146  | 
  finally have ri': "a \<otimes> (b \<otimes> i) = \<one>" .
  | 
| 
 | 
   147  | 
  | 
| 
 | 
   148  | 
  from carr' li' ri'
  | 
| 
 | 
   149  | 
      show "a \<in> Units G" by (simp add: Units_def, fast)
  | 
| 
 | 
   150  | 
qed
  | 
| 
 | 
   151  | 
  | 
| 
 | 
   152  | 
subsection {* Divisibility and association *}
 | 
| 
 | 
   153  | 
  | 
| 
 | 
   154  | 
subsubsection {* Function definitions *}
 | 
| 
 | 
   155  | 
  | 
| 
 | 
   156  | 
constdefs (structure G)
  | 
| 
 | 
   157  | 
  factor :: "[_, 'a, 'a] \<Rightarrow> bool" (infix "divides\<index>" 65)
  | 
| 
 | 
   158  | 
  "a divides b == \<exists>c\<in>carrier G. b = a \<otimes> c"
  | 
| 
 | 
   159  | 
  | 
| 
 | 
   160  | 
constdefs (structure G)
  | 
| 
 | 
   161  | 
  associated :: "[_, 'a, 'a] => bool" (infix "\<sim>\<index>" 55)
  | 
| 
 | 
   162  | 
  "a \<sim> b == a divides b \<and> b divides a"
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
abbreviation
  | 
| 
 | 
   165  | 
  "division_rel G == \<lparr>carrier = carrier G, eq = op \<sim>\<^bsub>G\<^esub>, le = op divides\<^bsub>G\<^esub>\<rparr>"
  | 
| 
 | 
   166  | 
  | 
| 
 | 
   167  | 
constdefs (structure G)
  | 
| 
 | 
   168  | 
  properfactor :: "[_, 'a, 'a] \<Rightarrow> bool"
  | 
| 
 | 
   169  | 
  "properfactor G a b == a divides b \<and> \<not>(b divides a)"
  | 
| 
 | 
   170  | 
  | 
| 
 | 
   171  | 
constdefs (structure G)
  | 
| 
 | 
   172  | 
  irreducible :: "[_, 'a] \<Rightarrow> bool"
  | 
| 
 | 
   173  | 
  "irreducible G a == a \<notin> Units G \<and> (\<forall>b\<in>carrier G. properfactor G b a \<longrightarrow> b \<in> Units G)"
  | 
| 
 | 
   174  | 
  | 
| 
 | 
   175  | 
constdefs (structure G)
  | 
| 
 | 
   176  | 
  prime :: "[_, 'a] \<Rightarrow> bool"
  | 
| 
 | 
   177  | 
  "prime G p == p \<notin> Units G \<and> 
  | 
| 
 | 
   178  | 
                (\<forall>a\<in>carrier G. \<forall>b\<in>carrier G. p divides (a \<otimes> b) \<longrightarrow> p divides a \<or> p divides b)"
  | 
| 
 | 
   179  | 
  | 
| 
 | 
   180  | 
  | 
| 
 | 
   181  | 
  | 
| 
 | 
   182  | 
subsubsection {* Divisibility *}
 | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
lemma dividesI:
  | 
| 
 | 
   185  | 
  fixes G (structure)
  | 
| 
 | 
   186  | 
  assumes carr: "c \<in> carrier G"
  | 
| 
 | 
   187  | 
    and p: "b = a \<otimes> c"
  | 
| 
 | 
   188  | 
  shows "a divides b"
  | 
| 
 | 
   189  | 
unfolding factor_def
  | 
| 
 | 
   190  | 
using assms by fast
  | 
| 
 | 
   191  | 
  | 
| 
 | 
   192  | 
lemma dividesI' [intro]:
  | 
| 
 | 
   193  | 
   fixes G (structure)
  | 
| 
 | 
   194  | 
  assumes p: "b = a \<otimes> c"
  | 
| 
 | 
   195  | 
    and carr: "c \<in> carrier G"
  | 
| 
 | 
   196  | 
  shows "a divides b"
  | 
| 
 | 
   197  | 
using assms
  | 
| 
 | 
   198  | 
by (fast intro: dividesI)
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
lemma dividesD:
  | 
| 
 | 
   201  | 
  fixes G (structure)
  | 
| 
 | 
   202  | 
  assumes "a divides b"
  | 
| 
 | 
   203  | 
  shows "\<exists>c\<in>carrier G. b = a \<otimes> c"
  | 
| 
 | 
   204  | 
using assms
  | 
| 
 | 
   205  | 
unfolding factor_def
  | 
| 
 | 
   206  | 
by fast
  | 
| 
 | 
   207  | 
  | 
| 
 | 
   208  | 
lemma dividesE [elim]:
  | 
| 
 | 
   209  | 
  fixes G (structure)
  | 
| 
 | 
   210  | 
  assumes d: "a divides b"
  | 
| 
 | 
   211  | 
    and elim: "\<And>c. \<lbrakk>b = a \<otimes> c; c \<in> carrier G\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   212  | 
  shows "P"
  | 
| 
 | 
   213  | 
proof -
  | 
| 
 | 
   214  | 
  from dividesD[OF d]
  | 
| 
 | 
   215  | 
      obtain c
  | 
| 
 | 
   216  | 
      where "c\<in>carrier G"
  | 
| 
 | 
   217  | 
      and "b = a \<otimes> c"
  | 
| 
 | 
   218  | 
      by auto
  | 
| 
 | 
   219  | 
  thus "P" by (elim elim)
  | 
| 
 | 
   220  | 
qed
  | 
| 
 | 
   221  | 
  | 
| 
 | 
   222  | 
lemma (in monoid) divides_refl[simp, intro!]:
  | 
| 
 | 
   223  | 
  assumes carr: "a \<in> carrier G"
  | 
| 
 | 
   224  | 
  shows "a divides a"
  | 
| 
 | 
   225  | 
apply (intro dividesI[of "\<one>"])
  | 
| 
 | 
   226  | 
apply (simp, simp add: carr)
  | 
| 
 | 
   227  | 
done
  | 
| 
 | 
   228  | 
  | 
| 
 | 
   229  | 
lemma (in monoid) divides_trans [trans]:
  | 
| 
 | 
   230  | 
  assumes dvds: "a divides b"  "b divides c"
  | 
| 
 | 
   231  | 
    and acarr: "a \<in> carrier G"
  | 
| 
 | 
   232  | 
  shows "a divides c"
  | 
| 
 | 
   233  | 
using dvds[THEN dividesD]
  | 
| 
 | 
   234  | 
by (blast intro: dividesI m_assoc acarr)
  | 
| 
 | 
   235  | 
  | 
| 
 | 
   236  | 
lemma (in monoid) divides_mult_lI [intro]:
  | 
| 
 | 
   237  | 
  assumes ab: "a divides b"
  | 
| 
 | 
   238  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   239  | 
  shows "(c \<otimes> a) divides (c \<otimes> b)"
  | 
| 
 | 
   240  | 
using ab
  | 
| 
 | 
   241  | 
apply (elim dividesE, simp add: m_assoc[symmetric] carr)
  | 
| 
 | 
   242  | 
apply (fast intro: dividesI)
  | 
| 
 | 
   243  | 
done
  | 
| 
 | 
   244  | 
  | 
| 
 | 
   245  | 
lemma (in monoid_cancel) divides_mult_l [simp]:
  | 
| 
 | 
   246  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   247  | 
  shows "(c \<otimes> a) divides (c \<otimes> b) = a divides b"
  | 
| 
 | 
   248  | 
apply safe
  | 
| 
 | 
   249  | 
 apply (elim dividesE, intro dividesI, assumption)
  | 
| 
 | 
   250  | 
 apply (rule l_cancel[of c])
  | 
| 
 | 
   251  | 
    apply (simp add: m_assoc carr)+
  | 
| 
 | 
   252  | 
apply (fast intro: divides_mult_lI carr)
  | 
| 
 | 
   253  | 
done
  | 
| 
 | 
   254  | 
  | 
| 
 | 
   255  | 
lemma (in comm_monoid) divides_mult_rI [intro]:
  | 
| 
 | 
   256  | 
  assumes ab: "a divides b"
  | 
| 
 | 
   257  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   258  | 
  shows "(a \<otimes> c) divides (b \<otimes> c)"
  | 
| 
 | 
   259  | 
using carr ab
  | 
| 
 | 
   260  | 
apply (simp add: m_comm[of a c] m_comm[of b c])
  | 
| 
 | 
   261  | 
apply (rule divides_mult_lI, assumption+)
  | 
| 
 | 
   262  | 
done
  | 
| 
 | 
   263  | 
  | 
| 
 | 
   264  | 
lemma (in comm_monoid_cancel) divides_mult_r [simp]:
  | 
| 
 | 
   265  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   266  | 
  shows "(a \<otimes> c) divides (b \<otimes> c) = a divides b"
  | 
| 
 | 
   267  | 
using carr
  | 
| 
 | 
   268  | 
by (simp add: m_comm[of a c] m_comm[of b c])
  | 
| 
 | 
   269  | 
  | 
| 
 | 
   270  | 
lemma (in monoid) divides_prod_r:
  | 
| 
 | 
   271  | 
  assumes ab: "a divides b"
  | 
| 
 | 
   272  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   273  | 
  shows "a divides (b \<otimes> c)"
  | 
| 
 | 
   274  | 
using ab carr
  | 
| 
 | 
   275  | 
by (fast intro: m_assoc)
  | 
| 
 | 
   276  | 
  | 
| 
 | 
   277  | 
lemma (in comm_monoid) divides_prod_l:
  | 
| 
 | 
   278  | 
  assumes carr[intro]: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   279  | 
    and ab: "a divides b"
  | 
| 
 | 
   280  | 
  shows "a divides (c \<otimes> b)"
  | 
| 
 | 
   281  | 
using ab carr
  | 
| 
 | 
   282  | 
apply (simp add: m_comm[of c b])
  | 
| 
 | 
   283  | 
apply (fast intro: divides_prod_r)
  | 
| 
 | 
   284  | 
done
  | 
| 
 | 
   285  | 
  | 
| 
 | 
   286  | 
lemma (in monoid) unit_divides:
  | 
| 
 | 
   287  | 
  assumes uunit: "u \<in> Units G"
  | 
| 
 | 
   288  | 
      and acarr: "a \<in> carrier G"
  | 
| 
 | 
   289  | 
  shows "u divides a"
  | 
| 
 | 
   290  | 
proof (intro dividesI[of "(inv u) \<otimes> a"], fast intro: uunit acarr)
  | 
| 
 | 
   291  | 
  from uunit acarr
  | 
| 
 | 
   292  | 
      have xcarr: "inv u \<otimes> a \<in> carrier G" by fast
  | 
| 
 | 
   293  | 
  | 
| 
 | 
   294  | 
  from uunit acarr
  | 
| 
 | 
   295  | 
       have "u \<otimes> (inv u \<otimes> a) = (u \<otimes> inv u) \<otimes> a" by (fast intro: m_assoc[symmetric])
  | 
| 
 | 
   296  | 
  also have "\<dots> = \<one> \<otimes> a" by (simp add: Units_r_inv[OF uunit])
  | 
| 
 | 
   297  | 
  also from acarr 
  | 
| 
 | 
   298  | 
       have "\<dots> = a" by simp
  | 
| 
 | 
   299  | 
  finally
  | 
| 
 | 
   300  | 
       show "a = u \<otimes> (inv u \<otimes> a)" ..
  | 
| 
 | 
   301  | 
qed
  | 
| 
 | 
   302  | 
  | 
| 
 | 
   303  | 
lemma (in comm_monoid) divides_unit:
  | 
| 
 | 
   304  | 
  assumes udvd: "a divides u"
  | 
| 
 | 
   305  | 
      and  carr: "a \<in> carrier G"  "u \<in> Units G"
  | 
| 
 | 
   306  | 
  shows "a \<in> Units G"
  | 
| 
 | 
   307  | 
using udvd carr
  | 
| 
 | 
   308  | 
by (blast intro: unit_factor)
  | 
| 
 | 
   309  | 
  | 
| 
 | 
   310  | 
lemma (in comm_monoid) Unit_eq_dividesone:
  | 
| 
 | 
   311  | 
  assumes ucarr: "u \<in> carrier G"
  | 
| 
 | 
   312  | 
  shows "u \<in> Units G = u divides \<one>"
  | 
| 
 | 
   313  | 
using ucarr
  | 
| 
 | 
   314  | 
by (fast dest: divides_unit intro: unit_divides)
  | 
| 
 | 
   315  | 
  | 
| 
 | 
   316  | 
  | 
| 
 | 
   317  | 
subsubsection {* Association *}
 | 
| 
 | 
   318  | 
  | 
| 
 | 
   319  | 
lemma associatedI:
  | 
| 
 | 
   320  | 
  fixes G (structure)
  | 
| 
 | 
   321  | 
  assumes "a divides b"  "b divides a"
  | 
| 
 | 
   322  | 
  shows "a \<sim> b"
  | 
| 
 | 
   323  | 
using assms
  | 
| 
 | 
   324  | 
by (simp add: associated_def)
  | 
| 
 | 
   325  | 
  | 
| 
 | 
   326  | 
lemma (in monoid) associatedI2:
  | 
| 
 | 
   327  | 
  assumes uunit[simp]: "u \<in> Units G"
  | 
| 
 | 
   328  | 
    and a: "a = b \<otimes> u"
  | 
| 
 | 
   329  | 
    and bcarr[simp]: "b \<in> carrier G"
  | 
| 
 | 
   330  | 
  shows "a \<sim> b"
  | 
| 
 | 
   331  | 
using uunit bcarr
  | 
| 
 | 
   332  | 
unfolding a
  | 
| 
 | 
   333  | 
apply (intro associatedI)
  | 
| 
 | 
   334  | 
 apply (rule dividesI[of "inv u"], simp)
  | 
| 
 | 
   335  | 
 apply (simp add: m_assoc Units_closed Units_r_inv)
  | 
| 
 | 
   336  | 
apply fast
  | 
| 
 | 
   337  | 
done
  | 
| 
 | 
   338  | 
  | 
| 
 | 
   339  | 
lemma (in monoid) associatedI2':
  | 
| 
 | 
   340  | 
  assumes a: "a = b \<otimes> u"
  | 
| 
 | 
   341  | 
    and uunit: "u \<in> Units G"
  | 
| 
 | 
   342  | 
    and bcarr: "b \<in> carrier G"
  | 
| 
 | 
   343  | 
  shows "a \<sim> b"
  | 
| 
 | 
   344  | 
using assms by (intro associatedI2)
  | 
| 
 | 
   345  | 
  | 
| 
 | 
   346  | 
lemma associatedD:
  | 
| 
 | 
   347  | 
  fixes G (structure)
  | 
| 
 | 
   348  | 
  assumes "a \<sim> b"
  | 
| 
 | 
   349  | 
  shows "a divides b"
  | 
| 
 | 
   350  | 
using assms by (simp add: associated_def)
  | 
| 
 | 
   351  | 
  | 
| 
 | 
   352  | 
lemma (in monoid_cancel) associatedD2:
  | 
| 
 | 
   353  | 
  assumes assoc: "a \<sim> b"
  | 
| 
 | 
   354  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   355  | 
  shows "\<exists>u\<in>Units G. a = b \<otimes> u"
  | 
| 
 | 
   356  | 
using assoc
  | 
| 
 | 
   357  | 
unfolding associated_def
  | 
| 
 | 
   358  | 
proof clarify
  | 
| 
 | 
   359  | 
  assume "b divides a"
  | 
| 
 | 
   360  | 
  hence "\<exists>u\<in>carrier G. a = b \<otimes> u" by (rule dividesD)
  | 
| 
 | 
   361  | 
  from this obtain u
  | 
| 
 | 
   362  | 
      where ucarr: "u \<in> carrier G" and a: "a = b \<otimes> u"
  | 
| 
 | 
   363  | 
      by auto
  | 
| 
 | 
   364  | 
  | 
| 
 | 
   365  | 
  assume "a divides b"
  | 
| 
 | 
   366  | 
  hence "\<exists>u'\<in>carrier G. b = a \<otimes> u'" by (rule dividesD)
  | 
| 
 | 
   367  | 
  from this obtain u'
  | 
| 
 | 
   368  | 
      where u'carr: "u' \<in> carrier G" and b: "b = a \<otimes> u'"
  | 
| 
 | 
   369  | 
      by auto
  | 
| 
 | 
   370  | 
  note carr = carr ucarr u'carr
  | 
| 
 | 
   371  | 
  | 
| 
 | 
   372  | 
  from carr
  | 
| 
 | 
   373  | 
       have "a \<otimes> \<one> = a" by simp
  | 
| 
 | 
   374  | 
  also have "\<dots> = b \<otimes> u" by (simp add: a)
  | 
| 
 | 
   375  | 
  also have "\<dots> = a \<otimes> u' \<otimes> u" by (simp add: b)
  | 
| 
 | 
   376  | 
  also from carr
  | 
| 
 | 
   377  | 
       have "\<dots> = a \<otimes> (u' \<otimes> u)" by (simp add: m_assoc)
  | 
| 
 | 
   378  | 
  finally
  | 
| 
 | 
   379  | 
       have "a \<otimes> \<one> = a \<otimes> (u' \<otimes> u)" .
  | 
| 
 | 
   380  | 
  with carr
  | 
| 
 | 
   381  | 
      have u1: "\<one> = u' \<otimes> u" by (fast dest: l_cancel)
  | 
| 
 | 
   382  | 
  | 
| 
 | 
   383  | 
  from carr
  | 
| 
 | 
   384  | 
       have "b \<otimes> \<one> = b" by simp
  | 
| 
 | 
   385  | 
  also have "\<dots> = a \<otimes> u'" by (simp add: b)
  | 
| 
 | 
   386  | 
  also have "\<dots> = b \<otimes> u \<otimes> u'" by (simp add: a)
  | 
| 
 | 
   387  | 
  also from carr
  | 
| 
 | 
   388  | 
       have "\<dots> = b \<otimes> (u \<otimes> u')" by (simp add: m_assoc)
  | 
| 
 | 
   389  | 
  finally
  | 
| 
 | 
   390  | 
       have "b \<otimes> \<one> = b \<otimes> (u \<otimes> u')" .
  | 
| 
 | 
   391  | 
  with carr
  | 
| 
 | 
   392  | 
      have u2: "\<one> = u \<otimes> u'" by (fast dest: l_cancel)
  | 
| 
 | 
   393  | 
  | 
| 
 | 
   394  | 
  from u'carr u1[symmetric] u2[symmetric]
  | 
| 
 | 
   395  | 
      have "\<exists>u'\<in>carrier G. u' \<otimes> u = \<one> \<and> u \<otimes> u' = \<one>" by fast
  | 
| 
 | 
   396  | 
  hence "u \<in> Units G" by (simp add: Units_def ucarr)
  | 
| 
 | 
   397  | 
  | 
| 
 | 
   398  | 
  from ucarr this a
  | 
| 
 | 
   399  | 
      show "\<exists>u\<in>Units G. a = b \<otimes> u" by fast
  | 
| 
 | 
   400  | 
qed
  | 
| 
 | 
   401  | 
  | 
| 
 | 
   402  | 
lemma associatedE:
  | 
| 
 | 
   403  | 
  fixes G (structure)
  | 
| 
 | 
   404  | 
  assumes assoc: "a \<sim> b"
  | 
| 
 | 
   405  | 
    and e: "\<lbrakk>a divides b; b divides a\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   406  | 
  shows "P"
  | 
| 
 | 
   407  | 
proof -
  | 
| 
 | 
   408  | 
  from assoc
  | 
| 
 | 
   409  | 
      have "a divides b"  "b divides a"
  | 
| 
 | 
   410  | 
      by (simp add: associated_def)+
  | 
| 
 | 
   411  | 
  thus "P" by (elim e)
  | 
| 
 | 
   412  | 
qed
  | 
| 
 | 
   413  | 
  | 
| 
 | 
   414  | 
lemma (in monoid_cancel) associatedE2:
  | 
| 
 | 
   415  | 
  assumes assoc: "a \<sim> b"
  | 
| 
 | 
   416  | 
    and e: "\<And>u. \<lbrakk>a = b \<otimes> u; u \<in> Units G\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   417  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   418  | 
  shows "P"
  | 
| 
 | 
   419  | 
proof -
  | 
| 
 | 
   420  | 
  from assoc and carr
  | 
| 
 | 
   421  | 
      have "\<exists>u\<in>Units G. a = b \<otimes> u" by (rule associatedD2)
  | 
| 
 | 
   422  | 
  from this obtain u
  | 
| 
 | 
   423  | 
      where "u \<in> Units G"  "a = b \<otimes> u"
  | 
| 
 | 
   424  | 
      by auto
  | 
| 
 | 
   425  | 
  thus "P" by (elim e)
  | 
| 
 | 
   426  | 
qed
  | 
| 
 | 
   427  | 
  | 
| 
 | 
   428  | 
lemma (in monoid) associated_refl [simp, intro!]:
  | 
| 
 | 
   429  | 
  assumes "a \<in> carrier G"
  | 
| 
 | 
   430  | 
  shows "a \<sim> a"
  | 
| 
 | 
   431  | 
using assms
  | 
| 
 | 
   432  | 
by (fast intro: associatedI)
  | 
| 
 | 
   433  | 
  | 
| 
 | 
   434  | 
lemma (in monoid) associated_sym [sym]:
  | 
| 
 | 
   435  | 
  assumes "a \<sim> b"
  | 
| 
 | 
   436  | 
    and "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   437  | 
  shows "b \<sim> a"
  | 
| 
 | 
   438  | 
using assms
  | 
| 
 | 
   439  | 
by (iprover intro: associatedI elim: associatedE)
  | 
| 
 | 
   440  | 
  | 
| 
 | 
   441  | 
lemma (in monoid) associated_trans [trans]:
  | 
| 
 | 
   442  | 
  assumes "a \<sim> b"  "b \<sim> c"
  | 
| 
 | 
   443  | 
    and "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   444  | 
  shows "a \<sim> c"
  | 
| 
 | 
   445  | 
using assms
  | 
| 
 | 
   446  | 
by (iprover intro: associatedI divides_trans elim: associatedE)
  | 
| 
 | 
   447  | 
  | 
| 
 | 
   448  | 
lemma (in monoid) division_equiv [intro, simp]:
  | 
| 
 | 
   449  | 
  "equivalence (division_rel G)"
  | 
| 
 | 
   450  | 
  apply unfold_locales
  | 
| 
 | 
   451  | 
  apply simp_all
  | 
| 
 | 
   452  | 
  apply (rule associated_sym, assumption+)
  | 
| 
 | 
   453  | 
  apply (iprover intro: associated_trans)
  | 
| 
 | 
   454  | 
  done
  | 
| 
 | 
   455  | 
  | 
| 
 | 
   456  | 
  | 
| 
 | 
   457  | 
subsubsection {* Division and associativity *}
 | 
| 
 | 
   458  | 
  | 
| 
 | 
   459  | 
lemma divides_antisym:
  | 
| 
 | 
   460  | 
  fixes G (structure)
  | 
| 
 | 
   461  | 
  assumes "a divides b"  "b divides a"
  | 
| 
 | 
   462  | 
    and "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   463  | 
  shows "a \<sim> b"
  | 
| 
 | 
   464  | 
using assms
  | 
| 
 | 
   465  | 
by (fast intro: associatedI)
  | 
| 
 | 
   466  | 
  | 
| 
 | 
   467  | 
lemma (in monoid) divides_cong_l [trans]:
  | 
| 
 | 
   468  | 
  assumes xx': "x \<sim> x'"
  | 
| 
 | 
   469  | 
    and xdvdy: "x' divides y"
  | 
| 
 | 
   470  | 
    and carr [simp]: "x \<in> carrier G"  "x' \<in> carrier G"  "y \<in> carrier G"
  | 
| 
 | 
   471  | 
  shows "x divides y"
  | 
| 
 | 
   472  | 
proof -
  | 
| 
 | 
   473  | 
  from xx'
  | 
| 
 | 
   474  | 
       have "x divides x'" by (simp add: associatedD)
  | 
| 
 | 
   475  | 
  also note xdvdy
  | 
| 
 | 
   476  | 
  finally
  | 
| 
 | 
   477  | 
       show "x divides y" by simp
  | 
| 
 | 
   478  | 
qed
  | 
| 
 | 
   479  | 
  | 
| 
 | 
   480  | 
lemma (in monoid) divides_cong_r [trans]:
  | 
| 
 | 
   481  | 
  assumes xdvdy: "x divides y"
  | 
| 
 | 
   482  | 
    and yy': "y \<sim> y'"
  | 
| 
 | 
   483  | 
    and carr[simp]: "x \<in> carrier G"  "y \<in> carrier G"  "y' \<in> carrier G"
  | 
| 
 | 
   484  | 
  shows "x divides y'"
  | 
| 
 | 
   485  | 
proof -
  | 
| 
 | 
   486  | 
  note xdvdy
  | 
| 
 | 
   487  | 
  also from yy'
  | 
| 
 | 
   488  | 
       have "y divides y'" by (simp add: associatedD)
  | 
| 
 | 
   489  | 
  finally
  | 
| 
 | 
   490  | 
       show "x divides y'" by simp
  | 
| 
 | 
   491  | 
qed
  | 
| 
 | 
   492  | 
  | 
| 
 | 
   493  | 
lemma (in monoid) division_gpartial_order [simp, intro!]:
  | 
| 
 | 
   494  | 
  "gpartial_order (division_rel G)"
  | 
| 
 | 
   495  | 
  apply unfold_locales
  | 
| 
 | 
   496  | 
  apply simp_all
  | 
| 
 | 
   497  | 
  apply (simp add: associated_sym)
  | 
| 
 | 
   498  | 
  apply (blast intro: associated_trans)
  | 
| 
 | 
   499  | 
  apply (simp add: divides_antisym)
  | 
| 
 | 
   500  | 
  apply (blast intro: divides_trans)
  | 
| 
 | 
   501  | 
  apply (blast intro: divides_cong_l divides_cong_r associated_sym)
  | 
| 
 | 
   502  | 
  done
  | 
| 
 | 
   503  | 
  | 
| 
 | 
   504  | 
    
  | 
| 
 | 
   505  | 
subsubsection {* Multiplication and associativity *}
 | 
| 
 | 
   506  | 
  | 
| 
 | 
   507  | 
lemma (in monoid_cancel) mult_cong_r:
  | 
| 
 | 
   508  | 
  assumes "b \<sim> b'"
  | 
| 
 | 
   509  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "b' \<in> carrier G"
  | 
| 
 | 
   510  | 
  shows "a \<otimes> b \<sim> a \<otimes> b'"
  | 
| 
 | 
   511  | 
using assms
  | 
| 
 | 
   512  | 
apply (elim associatedE2, intro associatedI2)
  | 
| 
 | 
   513  | 
apply (auto intro: m_assoc[symmetric])
  | 
| 
 | 
   514  | 
done
  | 
| 
 | 
   515  | 
  | 
| 
 | 
   516  | 
lemma (in comm_monoid_cancel) mult_cong_l:
  | 
| 
 | 
   517  | 
  assumes "a \<sim> a'"
  | 
| 
 | 
   518  | 
    and carr: "a \<in> carrier G"  "a' \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   519  | 
  shows "a \<otimes> b \<sim> a' \<otimes> b"
  | 
| 
 | 
   520  | 
using assms
  | 
| 
 | 
   521  | 
apply (elim associatedE2, intro associatedI2)
  | 
| 
 | 
   522  | 
    apply assumption
  | 
| 
 | 
   523  | 
   apply (simp add: m_assoc Units_closed)
  | 
| 
 | 
   524  | 
   apply (simp add: m_comm Units_closed)
  | 
| 
 | 
   525  | 
  apply simp+
  | 
| 
 | 
   526  | 
done
  | 
| 
 | 
   527  | 
  | 
| 
 | 
   528  | 
lemma (in monoid_cancel) assoc_l_cancel:
  | 
| 
 | 
   529  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "b' \<in> carrier G"
  | 
| 
 | 
   530  | 
    and "a \<otimes> b \<sim> a \<otimes> b'"
  | 
| 
 | 
   531  | 
  shows "b \<sim> b'"
  | 
| 
 | 
   532  | 
using assms
  | 
| 
 | 
   533  | 
apply (elim associatedE2, intro associatedI2)
  | 
| 
 | 
   534  | 
    apply assumption
  | 
| 
 | 
   535  | 
   apply (rule l_cancel[of a])
  | 
| 
 | 
   536  | 
      apply (simp add: m_assoc Units_closed)
  | 
| 
 | 
   537  | 
     apply fast+
  | 
| 
 | 
   538  | 
done
  | 
| 
 | 
   539  | 
  | 
| 
 | 
   540  | 
lemma (in comm_monoid_cancel) assoc_r_cancel:
  | 
| 
 | 
   541  | 
  assumes "a \<otimes> b \<sim> a' \<otimes> b"
  | 
| 
 | 
   542  | 
    and carr: "a \<in> carrier G"  "a' \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   543  | 
  shows "a \<sim> a'"
  | 
| 
 | 
   544  | 
using assms
  | 
| 
 | 
   545  | 
apply (elim associatedE2, intro associatedI2)
  | 
| 
 | 
   546  | 
    apply assumption
  | 
| 
 | 
   547  | 
   apply (rule r_cancel[of a b])
  | 
| 
 | 
   548  | 
      apply (simp add: m_assoc Units_closed)
  | 
| 
 | 
   549  | 
      apply (simp add: m_comm Units_closed)
  | 
| 
 | 
   550  | 
     apply fast+
  | 
| 
 | 
   551  | 
done
  | 
| 
 | 
   552  | 
  | 
| 
 | 
   553  | 
  | 
| 
 | 
   554  | 
subsubsection {* Units *}
 | 
| 
 | 
   555  | 
  | 
| 
 | 
   556  | 
lemma (in monoid_cancel) assoc_unit_l [trans]:
  | 
| 
 | 
   557  | 
  assumes asc: "a \<sim> b" and bunit: "b \<in> Units G"
  | 
| 
 | 
   558  | 
    and carr: "a \<in> carrier G" 
  | 
| 
 | 
   559  | 
  shows "a \<in> Units G"
  | 
| 
 | 
   560  | 
using assms
  | 
| 
 | 
   561  | 
by (fast elim: associatedE2)
  | 
| 
 | 
   562  | 
  | 
| 
 | 
   563  | 
lemma (in monoid_cancel) assoc_unit_r [trans]:
  | 
| 
 | 
   564  | 
  assumes aunit: "a \<in> Units G" and asc: "a \<sim> b"
  | 
| 
 | 
   565  | 
    and bcarr: "b \<in> carrier G"
  | 
| 
 | 
   566  | 
  shows "b \<in> Units G"
  | 
| 
 | 
   567  | 
using aunit bcarr associated_sym[OF asc]
  | 
| 
 | 
   568  | 
by (blast intro: assoc_unit_l)
  | 
| 
 | 
   569  | 
  | 
| 
 | 
   570  | 
lemma (in comm_monoid) Units_cong:
  | 
| 
 | 
   571  | 
  assumes aunit: "a \<in> Units G" and asc: "a \<sim> b"
  | 
| 
 | 
   572  | 
    and bcarr: "b \<in> carrier G"
  | 
| 
 | 
   573  | 
  shows "b \<in> Units G"
  | 
| 
 | 
   574  | 
using assms
  | 
| 
 | 
   575  | 
by (blast intro: divides_unit elim: associatedE)
  | 
| 
 | 
   576  | 
  | 
| 
 | 
   577  | 
lemma (in monoid) Units_assoc:
  | 
| 
 | 
   578  | 
  assumes units: "a \<in> Units G"  "b \<in> Units G"
  | 
| 
 | 
   579  | 
  shows "a \<sim> b"
  | 
| 
 | 
   580  | 
using units
  | 
| 
 | 
   581  | 
by (fast intro: associatedI unit_divides)
  | 
| 
 | 
   582  | 
  | 
| 
 | 
   583  | 
lemma (in monoid) Units_are_ones:
  | 
| 
 | 
   584  | 
  "Units G {.=}\<^bsub>(division_rel G)\<^esub> {\<one>}"
 | 
| 
 | 
   585  | 
apply (simp add: set_eq_def elem_def, rule, simp_all)
  | 
| 
 | 
   586  | 
proof clarsimp
  | 
| 
 | 
   587  | 
  fix a
  | 
| 
 | 
   588  | 
  assume aunit: "a \<in> Units G"
  | 
| 
 | 
   589  | 
  show "a \<sim> \<one>"
  | 
| 
 | 
   590  | 
  apply (rule associatedI)
  | 
| 
 | 
   591  | 
   apply (fast intro: dividesI[of "inv a"] aunit Units_r_inv[symmetric])
  | 
| 
 | 
   592  | 
  apply (fast intro: dividesI[of "a"] l_one[symmetric] Units_closed[OF aunit])
  | 
| 
 | 
   593  | 
  done
  | 
| 
 | 
   594  | 
next
  | 
| 
 | 
   595  | 
  have "\<one> \<in> Units G" by simp
  | 
| 
 | 
   596  | 
  moreover have "\<one> \<sim> \<one>" by simp
  | 
| 
 | 
   597  | 
  ultimately show "\<exists>a \<in> Units G. \<one> \<sim> a" by fast
  | 
| 
 | 
   598  | 
qed
  | 
| 
 | 
   599  | 
  | 
| 
 | 
   600  | 
lemma (in comm_monoid) Units_Lower:
  | 
| 
 | 
   601  | 
  "Units G = Lower (division_rel G) (carrier G)"
  | 
| 
 | 
   602  | 
apply (simp add: Units_def Lower_def)
  | 
| 
 | 
   603  | 
apply (rule, rule)
  | 
| 
 | 
   604  | 
 apply clarsimp
  | 
| 
 | 
   605  | 
  apply (rule unit_divides)
  | 
| 
 | 
   606  | 
   apply (unfold Units_def, fast)
  | 
| 
 | 
   607  | 
  apply assumption
  | 
| 
 | 
   608  | 
apply clarsimp
  | 
| 
 | 
   609  | 
proof -
  | 
| 
 | 
   610  | 
  fix x
  | 
| 
 | 
   611  | 
  assume xcarr: "x \<in> carrier G"
  | 
| 
 | 
   612  | 
  assume r[rule_format]: "\<forall>y. y \<in> carrier G \<longrightarrow> x divides y"
  | 
| 
 | 
   613  | 
  have "\<one> \<in> carrier G" by simp
  | 
| 
 | 
   614  | 
  hence "x divides \<one>" by (rule r)
  | 
| 
 | 
   615  | 
  hence "\<exists>x'\<in>carrier G. \<one> = x \<otimes> x'" by (rule dividesE, fast)
  | 
| 
 | 
   616  | 
  from this obtain x'
  | 
| 
 | 
   617  | 
      where x'carr: "x' \<in> carrier G"
  | 
| 
 | 
   618  | 
      and xx': "\<one> = x \<otimes> x'"
  | 
| 
 | 
   619  | 
      by auto
  | 
| 
 | 
   620  | 
  | 
| 
 | 
   621  | 
  note xx'
  | 
| 
 | 
   622  | 
  also with xcarr x'carr
  | 
| 
 | 
   623  | 
       have "\<dots> = x' \<otimes> x" by (simp add: m_comm)
  | 
| 
 | 
   624  | 
  finally
  | 
| 
 | 
   625  | 
       have "\<one> = x' \<otimes> x" .
  | 
| 
 | 
   626  | 
  | 
| 
 | 
   627  | 
  from x'carr xx'[symmetric] this[symmetric]
  | 
| 
 | 
   628  | 
      show "\<exists>y\<in>carrier G. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast
  | 
| 
 | 
   629  | 
qed
  | 
| 
 | 
   630  | 
  | 
| 
 | 
   631  | 
  | 
| 
 | 
   632  | 
subsubsection {* Proper factors *}
 | 
| 
 | 
   633  | 
  | 
| 
 | 
   634  | 
lemma properfactorI:
  | 
| 
 | 
   635  | 
  fixes G (structure)
  | 
| 
 | 
   636  | 
  assumes "a divides b"
  | 
| 
 | 
   637  | 
    and "\<not>(b divides a)"
  | 
| 
 | 
   638  | 
  shows "properfactor G a b"
  | 
| 
 | 
   639  | 
using assms
  | 
| 
 | 
   640  | 
unfolding properfactor_def
  | 
| 
 | 
   641  | 
by simp
  | 
| 
 | 
   642  | 
  | 
| 
 | 
   643  | 
lemma properfactorI2:
  | 
| 
 | 
   644  | 
  fixes G (structure)
  | 
| 
 | 
   645  | 
  assumes advdb: "a divides b"
  | 
| 
 | 
   646  | 
    and neq: "\<not>(a \<sim> b)"
  | 
| 
 | 
   647  | 
  shows "properfactor G a b"
  | 
| 
 | 
   648  | 
apply (rule properfactorI, rule advdb)
  | 
| 
 | 
   649  | 
proof (rule ccontr, simp)
  | 
| 
 | 
   650  | 
  assume "b divides a"
  | 
| 
 | 
   651  | 
  with advdb have "a \<sim> b" by (rule associatedI)
  | 
| 
 | 
   652  | 
  with neq show "False" by fast
  | 
| 
 | 
   653  | 
qed
  | 
| 
 | 
   654  | 
  | 
| 
 | 
   655  | 
lemma (in comm_monoid_cancel) properfactorI3:
  | 
| 
 | 
   656  | 
  assumes p: "p = a \<otimes> b"
  | 
| 
 | 
   657  | 
    and nunit: "b \<notin> Units G"
  | 
| 
 | 
   658  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "p \<in> carrier G"
  | 
| 
 | 
   659  | 
  shows "properfactor G a p"
  | 
| 
 | 
   660  | 
unfolding p
  | 
| 
 | 
   661  | 
using carr
  | 
| 
 | 
   662  | 
apply (intro properfactorI, fast)
  | 
| 
 | 
   663  | 
proof (clarsimp, elim dividesE)
  | 
| 
 | 
   664  | 
  fix c
  | 
| 
 | 
   665  | 
  assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
   666  | 
  note [simp] = carr ccarr
  | 
| 
 | 
   667  | 
  | 
| 
 | 
   668  | 
  have "a \<otimes> \<one> = a" by simp
  | 
| 
 | 
   669  | 
  also assume "a = a \<otimes> b \<otimes> c"
  | 
| 
 | 
   670  | 
  also have "\<dots> = a \<otimes> (b \<otimes> c)" by (simp add: m_assoc)
  | 
| 
 | 
   671  | 
  finally have "a \<otimes> \<one> = a \<otimes> (b \<otimes> c)" .
  | 
| 
 | 
   672  | 
  | 
| 
 | 
   673  | 
  hence rinv: "\<one> = b \<otimes> c" by (intro l_cancel[of "a" "\<one>" "b \<otimes> c"], simp+)
  | 
| 
 | 
   674  | 
  also have "\<dots> = c \<otimes> b" by (simp add: m_comm)
  | 
| 
 | 
   675  | 
  finally have linv: "\<one> = c \<otimes> b" .
  | 
| 
 | 
   676  | 
  | 
| 
 | 
   677  | 
  from ccarr linv[symmetric] rinv[symmetric]
  | 
| 
 | 
   678  | 
  have "b \<in> Units G" unfolding Units_def by fastsimp
  | 
| 
 | 
   679  | 
  with nunit
  | 
| 
 | 
   680  | 
      show "False" ..
  | 
| 
 | 
   681  | 
qed
  | 
| 
 | 
   682  | 
  | 
| 
 | 
   683  | 
lemma properfactorE:
  | 
| 
 | 
   684  | 
  fixes G (structure)
  | 
| 
 | 
   685  | 
  assumes pf: "properfactor G a b"
  | 
| 
 | 
   686  | 
    and r: "\<lbrakk>a divides b; \<not>(b divides a)\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   687  | 
  shows "P"
  | 
| 
 | 
   688  | 
using pf
  | 
| 
 | 
   689  | 
unfolding properfactor_def
  | 
| 
 | 
   690  | 
by (fast intro: r)
  | 
| 
 | 
   691  | 
  | 
| 
 | 
   692  | 
lemma properfactorE2:
  | 
| 
 | 
   693  | 
  fixes G (structure)
  | 
| 
 | 
   694  | 
  assumes pf: "properfactor G a b"
  | 
| 
 | 
   695  | 
    and elim: "\<lbrakk>a divides b; \<not>(a \<sim> b)\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   696  | 
  shows "P"
  | 
| 
 | 
   697  | 
using pf
  | 
| 
 | 
   698  | 
unfolding properfactor_def
  | 
| 
 | 
   699  | 
by (fast elim: elim associatedE)
  | 
| 
 | 
   700  | 
  | 
| 
 | 
   701  | 
lemma (in monoid) properfactor_unitE:
  | 
| 
 | 
   702  | 
  assumes uunit: "u \<in> Units G"
  | 
| 
 | 
   703  | 
    and pf: "properfactor G a u"
  | 
| 
 | 
   704  | 
    and acarr: "a \<in> carrier G"
  | 
| 
 | 
   705  | 
  shows "P"
  | 
| 
 | 
   706  | 
using pf unit_divides[OF uunit acarr]
  | 
| 
 | 
   707  | 
by (fast elim: properfactorE)
  | 
| 
 | 
   708  | 
  | 
| 
 | 
   709  | 
  | 
| 
 | 
   710  | 
lemma (in monoid) properfactor_divides:
  | 
| 
 | 
   711  | 
  assumes pf: "properfactor G a b"
  | 
| 
 | 
   712  | 
  shows "a divides b"
  | 
| 
 | 
   713  | 
using pf
  | 
| 
 | 
   714  | 
by (elim properfactorE)
  | 
| 
 | 
   715  | 
  | 
| 
 | 
   716  | 
lemma (in monoid) properfactor_trans1 [trans]:
  | 
| 
 | 
   717  | 
  assumes dvds: "a divides b"  "properfactor G b c"
  | 
| 
 | 
   718  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   719  | 
  shows "properfactor G a c"
  | 
| 
 | 
   720  | 
using dvds carr
  | 
| 
 | 
   721  | 
apply (elim properfactorE, intro properfactorI)
  | 
| 
 | 
   722  | 
 apply (iprover intro: divides_trans)+
  | 
| 
 | 
   723  | 
done
  | 
| 
 | 
   724  | 
  | 
| 
 | 
   725  | 
lemma (in monoid) properfactor_trans2 [trans]:
  | 
| 
 | 
   726  | 
  assumes dvds: "properfactor G a b"  "b divides c"
  | 
| 
 | 
   727  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   728  | 
  shows "properfactor G a c"
  | 
| 
 | 
   729  | 
using dvds carr
  | 
| 
 | 
   730  | 
apply (elim properfactorE, intro properfactorI)
  | 
| 
 | 
   731  | 
 apply (iprover intro: divides_trans)+
  | 
| 
 | 
   732  | 
done
  | 
| 
 | 
   733  | 
  | 
| 
 | 
   734  | 
lemma properfactor_glless:
  | 
| 
 | 
   735  | 
  fixes G (structure)
  | 
| 
 | 
   736  | 
  shows "properfactor G = glless (division_rel G)"
  | 
| 
 | 
   737  | 
apply (rule ext) apply (rule ext) apply rule
  | 
| 
 | 
   738  | 
 apply (fastsimp elim: properfactorE2 intro: gllessI)
  | 
| 
 | 
   739  | 
apply (fastsimp elim: gllessE intro: properfactorI2)
  | 
| 
 | 
   740  | 
done
  | 
| 
 | 
   741  | 
  | 
| 
 | 
   742  | 
lemma (in monoid) properfactor_cong_l [trans]:
  | 
| 
 | 
   743  | 
  assumes x'x: "x' \<sim> x"
  | 
| 
 | 
   744  | 
    and pf: "properfactor G x y"
  | 
| 
 | 
   745  | 
    and carr: "x \<in> carrier G"  "x' \<in> carrier G"  "y \<in> carrier G"
  | 
| 
 | 
   746  | 
  shows "properfactor G x' y"
  | 
| 
 | 
   747  | 
using pf
  | 
| 
 | 
   748  | 
unfolding properfactor_glless
  | 
| 
 | 
   749  | 
proof -
  | 
| 
 | 
   750  | 
  interpret gpartial_order ["division_rel G"] ..
  | 
| 
 | 
   751  | 
  from x'x
  | 
| 
 | 
   752  | 
       have "x' .=\<^bsub>division_rel G\<^esub> x" by simp
  | 
| 
 | 
   753  | 
  also assume "x \<sqsubset>\<^bsub>division_rel G\<^esub> y"
  | 
| 
 | 
   754  | 
  finally
  | 
| 
 | 
   755  | 
       show "x' \<sqsubset>\<^bsub>division_rel G\<^esub> y" by (simp add: carr)
  | 
| 
 | 
   756  | 
qed
  | 
| 
 | 
   757  | 
  | 
| 
 | 
   758  | 
lemma (in monoid) properfactor_cong_r [trans]:
  | 
| 
 | 
   759  | 
  assumes pf: "properfactor G x y"
  | 
| 
 | 
   760  | 
    and yy': "y \<sim> y'"
  | 
| 
 | 
   761  | 
    and carr: "x \<in> carrier G"  "y \<in> carrier G"  "y' \<in> carrier G"
  | 
| 
 | 
   762  | 
  shows "properfactor G x y'"
  | 
| 
 | 
   763  | 
using pf
  | 
| 
 | 
   764  | 
unfolding properfactor_glless
  | 
| 
 | 
   765  | 
proof -
  | 
| 
 | 
   766  | 
  interpret gpartial_order ["division_rel G"] ..
  | 
| 
 | 
   767  | 
  assume "x \<sqsubset>\<^bsub>division_rel G\<^esub> y"
  | 
| 
 | 
   768  | 
  also from yy'
  | 
| 
 | 
   769  | 
       have "y .=\<^bsub>division_rel G\<^esub> y'" by simp
  | 
| 
 | 
   770  | 
  finally
  | 
| 
 | 
   771  | 
       show "x \<sqsubset>\<^bsub>division_rel G\<^esub> y'" by (simp add: carr)
  | 
| 
 | 
   772  | 
qed
  | 
| 
 | 
   773  | 
  | 
| 
 | 
   774  | 
lemma (in monoid_cancel) properfactor_mult_lI [intro]:
  | 
| 
 | 
   775  | 
  assumes ab: "properfactor G a b"
  | 
| 
 | 
   776  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   777  | 
  shows "properfactor G (c \<otimes> a) (c \<otimes> b)"
  | 
| 
 | 
   778  | 
using ab carr
  | 
| 
 | 
   779  | 
by (fastsimp elim: properfactorE intro: properfactorI)
  | 
| 
 | 
   780  | 
  | 
| 
 | 
   781  | 
lemma (in monoid_cancel) properfactor_mult_l [simp]:
  | 
| 
 | 
   782  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   783  | 
  shows "properfactor G (c \<otimes> a) (c \<otimes> b) = properfactor G a b"
  | 
| 
 | 
   784  | 
using carr
  | 
| 
 | 
   785  | 
by (fastsimp elim: properfactorE intro: properfactorI)
  | 
| 
 | 
   786  | 
  | 
| 
 | 
   787  | 
lemma (in comm_monoid_cancel) properfactor_mult_rI [intro]:
  | 
| 
 | 
   788  | 
  assumes ab: "properfactor G a b"
  | 
| 
 | 
   789  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   790  | 
  shows "properfactor G (a \<otimes> c) (b \<otimes> c)"
  | 
| 
 | 
   791  | 
using ab carr
  | 
| 
 | 
   792  | 
by (fastsimp elim: properfactorE intro: properfactorI)
  | 
| 
 | 
   793  | 
  | 
| 
 | 
   794  | 
lemma (in comm_monoid_cancel) properfactor_mult_r [simp]:
  | 
| 
 | 
   795  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   796  | 
  shows "properfactor G (a \<otimes> c) (b \<otimes> c) = properfactor G a b"
  | 
| 
 | 
   797  | 
using carr
  | 
| 
 | 
   798  | 
by (fastsimp elim: properfactorE intro: properfactorI)
  | 
| 
 | 
   799  | 
  | 
| 
 | 
   800  | 
lemma (in monoid) properfactor_prod_r:
  | 
| 
 | 
   801  | 
  assumes ab: "properfactor G a b"
  | 
| 
 | 
   802  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   803  | 
  shows "properfactor G a (b \<otimes> c)"
  | 
| 
 | 
   804  | 
by (intro properfactor_trans2[OF ab] divides_prod_r, simp+)
  | 
| 
 | 
   805  | 
  | 
| 
 | 
   806  | 
lemma (in comm_monoid) properfactor_prod_l:
  | 
| 
 | 
   807  | 
  assumes ab: "properfactor G a b"
  | 
| 
 | 
   808  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
   809  | 
  shows "properfactor G a (c \<otimes> b)"
  | 
| 
 | 
   810  | 
by (intro properfactor_trans2[OF ab] divides_prod_l, simp+)
  | 
| 
 | 
   811  | 
  | 
| 
 | 
   812  | 
  | 
| 
 | 
   813  | 
subsection {* Irreducible elements and primes *}
 | 
| 
 | 
   814  | 
  | 
| 
 | 
   815  | 
subsubsection {* Irreducible elements *}
 | 
| 
 | 
   816  | 
  | 
| 
 | 
   817  | 
lemma irreducibleI:
  | 
| 
 | 
   818  | 
  fixes G (structure)
  | 
| 
 | 
   819  | 
  assumes "a \<notin> Units G"
  | 
| 
 | 
   820  | 
    and "\<And>b. \<lbrakk>b \<in> carrier G; properfactor G b a\<rbrakk> \<Longrightarrow> b \<in> Units G"
  | 
| 
 | 
   821  | 
  shows "irreducible G a"
  | 
| 
 | 
   822  | 
using assms 
  | 
| 
 | 
   823  | 
unfolding irreducible_def
  | 
| 
 | 
   824  | 
by blast
  | 
| 
 | 
   825  | 
  | 
| 
 | 
   826  | 
lemma irreducibleE:
  | 
| 
 | 
   827  | 
  fixes G (structure)
  | 
| 
 | 
   828  | 
  assumes irr: "irreducible G a"
  | 
| 
 | 
   829  | 
     and elim: "\<lbrakk>a \<notin> Units G; \<forall>b. b \<in> carrier G \<and> properfactor G b a \<longrightarrow> b \<in> Units G\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   830  | 
  shows "P"
  | 
| 
 | 
   831  | 
using assms
  | 
| 
 | 
   832  | 
unfolding irreducible_def
  | 
| 
 | 
   833  | 
by blast
  | 
| 
 | 
   834  | 
  | 
| 
 | 
   835  | 
lemma irreducibleD:
  | 
| 
 | 
   836  | 
  fixes G (structure)
  | 
| 
 | 
   837  | 
  assumes irr: "irreducible G a"
  | 
| 
 | 
   838  | 
     and pf: "properfactor G b a"
  | 
| 
 | 
   839  | 
     and bcarr: "b \<in> carrier G"
  | 
| 
 | 
   840  | 
  shows "b \<in> Units G"
  | 
| 
 | 
   841  | 
using assms
  | 
| 
 | 
   842  | 
by (fast elim: irreducibleE)
  | 
| 
 | 
   843  | 
  | 
| 
 | 
   844  | 
lemma (in monoid_cancel) irreducible_cong [trans]:
  | 
| 
 | 
   845  | 
  assumes irred: "irreducible G a"
  | 
| 
 | 
   846  | 
    and aa': "a \<sim> a'"
  | 
| 
 | 
   847  | 
    and carr[simp]: "a \<in> carrier G"  "a' \<in> carrier G"
  | 
| 
 | 
   848  | 
  shows "irreducible G a'"
  | 
| 
 | 
   849  | 
using assms
  | 
| 
 | 
   850  | 
apply (elim irreducibleE, intro irreducibleI)
  | 
| 
 | 
   851  | 
apply simp_all
  | 
| 
 | 
   852  | 
proof clarify
  | 
| 
 | 
   853  | 
  assume "a' \<in> Units G"
  | 
| 
 | 
   854  | 
  also note aa'[symmetric]
  | 
| 
 | 
   855  | 
  finally have aunit: "a \<in> Units G" by simp
  | 
| 
 | 
   856  | 
  | 
| 
 | 
   857  | 
  assume "a \<notin> Units G"
  | 
| 
 | 
   858  | 
  with aunit
  | 
| 
 | 
   859  | 
      show "False" by fast
  | 
| 
 | 
   860  | 
next
  | 
| 
 | 
   861  | 
  fix b
  | 
| 
 | 
   862  | 
  assume r[rule_format]: "\<forall>b. b \<in> carrier G \<and> properfactor G b a \<longrightarrow> b \<in> Units G"
  | 
| 
 | 
   863  | 
    and bcarr[simp]: "b \<in> carrier G"
  | 
| 
 | 
   864  | 
  assume "properfactor G b a'"
  | 
| 
 | 
   865  | 
  also note aa'[symmetric]
  | 
| 
 | 
   866  | 
  finally
  | 
| 
 | 
   867  | 
       have "properfactor G b a" by simp
  | 
| 
 | 
   868  | 
  | 
| 
 | 
   869  | 
  with bcarr
  | 
| 
 | 
   870  | 
     show "b \<in> Units G" by (fast intro: r)
  | 
| 
 | 
   871  | 
qed
  | 
| 
 | 
   872  | 
  | 
| 
 | 
   873  | 
  | 
| 
 | 
   874  | 
lemma (in monoid) irreducible_prod_rI:
  | 
| 
 | 
   875  | 
  assumes airr: "irreducible G a"
  | 
| 
 | 
   876  | 
    and bunit: "b \<in> Units G"
  | 
| 
 | 
   877  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   878  | 
  shows "irreducible G (a \<otimes> b)"
  | 
| 
 | 
   879  | 
using airr carr bunit
  | 
| 
 | 
   880  | 
apply (elim irreducibleE, intro irreducibleI, clarify)
  | 
| 
 | 
   881  | 
 apply (subgoal_tac "a \<in> Units G", simp)
  | 
| 
 | 
   882  | 
 apply (intro prod_unit_r[of a b] carr bunit, assumption)
  | 
| 
 | 
   883  | 
proof -
  | 
| 
 | 
   884  | 
  fix c
  | 
| 
 | 
   885  | 
  assume [simp]: "c \<in> carrier G"
  | 
| 
 | 
   886  | 
    and r[rule_format]: "\<forall>b. b \<in> carrier G \<and> properfactor G b a \<longrightarrow> b \<in> Units G"
  | 
| 
 | 
   887  | 
  assume "properfactor G c (a \<otimes> b)"
  | 
| 
 | 
   888  | 
  also have "a \<otimes> b \<sim> a" by (intro associatedI2[OF bunit], simp+)
  | 
| 
 | 
   889  | 
  finally
  | 
| 
 | 
   890  | 
       have pfa: "properfactor G c a" by simp
  | 
| 
 | 
   891  | 
  show "c \<in> Units G" by (rule r, simp add: pfa)
  | 
| 
 | 
   892  | 
qed
  | 
| 
 | 
   893  | 
  | 
| 
 | 
   894  | 
lemma (in comm_monoid) irreducible_prod_lI:
  | 
| 
 | 
   895  | 
  assumes birr: "irreducible G b"
  | 
| 
 | 
   896  | 
    and aunit: "a \<in> Units G"
  | 
| 
 | 
   897  | 
    and carr [simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   898  | 
  shows "irreducible G (a \<otimes> b)"
  | 
| 
 | 
   899  | 
apply (subst m_comm, simp+)
  | 
| 
 | 
   900  | 
apply (intro irreducible_prod_rI assms)
  | 
| 
 | 
   901  | 
done
  | 
| 
 | 
   902  | 
  | 
| 
 | 
   903  | 
lemma (in comm_monoid_cancel) irreducible_prodE [elim]:
  | 
| 
 | 
   904  | 
  assumes irr: "irreducible G (a \<otimes> b)"
  | 
| 
 | 
   905  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   906  | 
    and e1: "\<lbrakk>irreducible G a; b \<in> Units G\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   907  | 
    and e2: "\<lbrakk>a \<in> Units G; irreducible G b\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   908  | 
  shows "P"
  | 
| 
 | 
   909  | 
using irr
  | 
| 
 | 
   910  | 
proof (elim irreducibleE)
  | 
| 
 | 
   911  | 
  assume abnunit: "a \<otimes> b \<notin> Units G"
  | 
| 
 | 
   912  | 
    and isunit[rule_format]: "\<forall>ba. ba \<in> carrier G \<and> properfactor G ba (a \<otimes> b) \<longrightarrow> ba \<in> Units G"
  | 
| 
 | 
   913  | 
  | 
| 
 | 
   914  | 
  show "P"
  | 
| 
 | 
   915  | 
  proof (cases "a \<in> Units G")
  | 
| 
 | 
   916  | 
    assume aunit: "a \<in>  Units G"
  | 
| 
 | 
   917  | 
  | 
| 
 | 
   918  | 
    have "irreducible G b"
  | 
| 
 | 
   919  | 
    apply (rule irreducibleI)
  | 
| 
 | 
   920  | 
    proof (rule ccontr, simp)
  | 
| 
 | 
   921  | 
      assume "b \<in> Units G"
  | 
| 
 | 
   922  | 
      with aunit have "(a \<otimes> b) \<in> Units G" by fast
  | 
| 
 | 
   923  | 
      with abnunit show "False" ..
  | 
| 
 | 
   924  | 
    next
  | 
| 
 | 
   925  | 
      fix c
  | 
| 
 | 
   926  | 
      assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
   927  | 
        and "properfactor G c b"
  | 
| 
 | 
   928  | 
      hence "properfactor G c (a \<otimes> b)" by (simp add: properfactor_prod_l[of c b a])
  | 
| 
 | 
   929  | 
      from ccarr this show "c \<in> Units G" by (fast intro: isunit)
  | 
| 
 | 
   930  | 
    qed
  | 
| 
 | 
   931  | 
  | 
| 
 | 
   932  | 
    from aunit this show "P" by (rule e2)
  | 
| 
 | 
   933  | 
  next
  | 
| 
 | 
   934  | 
    assume anunit: "a \<notin> Units G"
  | 
| 
 | 
   935  | 
    with carr have "properfactor G b (b \<otimes> a)" by (fast intro: properfactorI3)
  | 
| 
 | 
   936  | 
    hence bf: "properfactor G b (a \<otimes> b)" by (subst m_comm[of a b], simp+)
  | 
| 
 | 
   937  | 
    hence bunit: "b \<in> Units G" by (intro isunit, simp)
  | 
| 
 | 
   938  | 
  | 
| 
 | 
   939  | 
    have "irreducible G a"
  | 
| 
 | 
   940  | 
    apply (rule irreducibleI)
  | 
| 
 | 
   941  | 
    proof (rule ccontr, simp)
  | 
| 
 | 
   942  | 
      assume "a \<in> Units G"
  | 
| 
 | 
   943  | 
      with bunit have "(a \<otimes> b) \<in> Units G" by fast
  | 
| 
 | 
   944  | 
      with abnunit show "False" ..
  | 
| 
 | 
   945  | 
    next
  | 
| 
 | 
   946  | 
      fix c
  | 
| 
 | 
   947  | 
      assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
   948  | 
        and "properfactor G c a"
  | 
| 
 | 
   949  | 
      hence "properfactor G c (a \<otimes> b)" by (simp add: properfactor_prod_r[of c a b])
  | 
| 
 | 
   950  | 
      from ccarr this show "c \<in> Units G" by (fast intro: isunit)
  | 
| 
 | 
   951  | 
    qed
  | 
| 
 | 
   952  | 
  | 
| 
 | 
   953  | 
    from this bunit show "P" by (rule e1)
  | 
| 
 | 
   954  | 
  qed
  | 
| 
 | 
   955  | 
qed
  | 
| 
 | 
   956  | 
  | 
| 
 | 
   957  | 
  | 
| 
 | 
   958  | 
subsubsection {* Prime elements *}
 | 
| 
 | 
   959  | 
  | 
| 
 | 
   960  | 
lemma primeI:
  | 
| 
 | 
   961  | 
  fixes G (structure)
  | 
| 
 | 
   962  | 
  assumes "p \<notin> Units G"
  | 
| 
 | 
   963  | 
    and "\<And>a b. \<lbrakk>a \<in> carrier G; b \<in> carrier G; p divides (a \<otimes> b)\<rbrakk> \<Longrightarrow> p divides a \<or> p divides b"
  | 
| 
 | 
   964  | 
  shows "prime G p"
  | 
| 
 | 
   965  | 
using assms
  | 
| 
 | 
   966  | 
unfolding prime_def
  | 
| 
 | 
   967  | 
by blast
  | 
| 
 | 
   968  | 
  | 
| 
 | 
   969  | 
lemma primeE:
  | 
| 
 | 
   970  | 
  fixes G (structure)
  | 
| 
 | 
   971  | 
  assumes pprime: "prime G p"
  | 
| 
 | 
   972  | 
    and e: "\<lbrakk>p \<notin> Units G; \<forall>a\<in>carrier G. \<forall>b\<in>carrier G.
  | 
| 
 | 
   973  | 
                          p divides a \<otimes> b \<longrightarrow> p divides a \<or> p divides b\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
   974  | 
  shows "P"
  | 
| 
 | 
   975  | 
using pprime
  | 
| 
 | 
   976  | 
unfolding prime_def
  | 
| 
 | 
   977  | 
by (blast dest: e)
  | 
| 
 | 
   978  | 
  | 
| 
 | 
   979  | 
lemma (in comm_monoid_cancel) prime_divides:
  | 
| 
 | 
   980  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
   981  | 
    and pprime: "prime G p"
  | 
| 
 | 
   982  | 
    and pdvd: "p divides a \<otimes> b"
  | 
| 
 | 
   983  | 
  shows "p divides a \<or> p divides b"
  | 
| 
 | 
   984  | 
using assms
  | 
| 
 | 
   985  | 
by (blast elim: primeE)
  | 
| 
 | 
   986  | 
  | 
| 
 | 
   987  | 
lemma (in monoid_cancel) prime_cong [trans]:
  | 
| 
 | 
   988  | 
  assumes pprime: "prime G p"
  | 
| 
 | 
   989  | 
    and pp': "p \<sim> p'"
  | 
| 
 | 
   990  | 
    and carr[simp]: "p \<in> carrier G"  "p' \<in> carrier G"
  | 
| 
 | 
   991  | 
  shows "prime G p'"
  | 
| 
 | 
   992  | 
using pprime
  | 
| 
 | 
   993  | 
apply (elim primeE, intro primeI)
  | 
| 
 | 
   994  | 
proof clarify
  | 
| 
 | 
   995  | 
  assume pnunit: "p \<notin> Units G"
  | 
| 
 | 
   996  | 
  assume "p' \<in> Units G"
  | 
| 
 | 
   997  | 
  also note pp'[symmetric]
  | 
| 
 | 
   998  | 
  finally
  | 
| 
 | 
   999  | 
       have "p \<in> Units G" by simp
  | 
| 
 | 
  1000  | 
  with pnunit
  | 
| 
 | 
  1001  | 
       show False ..
  | 
| 
 | 
  1002  | 
next
  | 
| 
 | 
  1003  | 
  fix a b
  | 
| 
 | 
  1004  | 
  assume r[rule_format]:
  | 
| 
 | 
  1005  | 
         "\<forall>a\<in>carrier G. \<forall>b\<in>carrier G. p divides a \<otimes> b \<longrightarrow> p divides a \<or> p divides b"
  | 
| 
 | 
  1006  | 
  assume p'dvd: "p' divides a \<otimes> b"
  | 
| 
 | 
  1007  | 
    and carr'[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  1008  | 
  | 
| 
 | 
  1009  | 
  note pp'
  | 
| 
 | 
  1010  | 
  also note p'dvd
  | 
| 
 | 
  1011  | 
  finally
  | 
| 
 | 
  1012  | 
       have "p divides a \<otimes> b" by simp
  | 
| 
 | 
  1013  | 
  hence "p divides a \<or> p divides b" by (intro r, simp+)
  | 
| 
 | 
  1014  | 
  moreover {
 | 
| 
 | 
  1015  | 
    note pp'[symmetric]
  | 
| 
 | 
  1016  | 
    also assume "p divides a"
  | 
| 
 | 
  1017  | 
    finally
  | 
| 
 | 
  1018  | 
         have "p' divides a" by simp
  | 
| 
 | 
  1019  | 
    hence "p' divides a \<or> p' divides b" by simp
  | 
| 
 | 
  1020  | 
  }
  | 
| 
 | 
  1021  | 
  moreover {
 | 
| 
 | 
  1022  | 
    note pp'[symmetric]
  | 
| 
 | 
  1023  | 
    also assume "p divides b"
  | 
| 
 | 
  1024  | 
    finally
  | 
| 
 | 
  1025  | 
         have "p' divides b" by simp
  | 
| 
 | 
  1026  | 
    hence "p' divides a \<or> p' divides b" by simp
  | 
| 
 | 
  1027  | 
  }
  | 
| 
 | 
  1028  | 
  ultimately
  | 
| 
 | 
  1029  | 
    show "p' divides a \<or> p' divides b" by fast
  | 
| 
 | 
  1030  | 
qed
  | 
| 
 | 
  1031  | 
  | 
| 
 | 
  1032  | 
  | 
| 
 | 
  1033  | 
subsection {* Factorization and factorial monoids *}
 | 
| 
 | 
  1034  | 
  | 
| 
 | 
  1035  | 
(*
  | 
| 
 | 
  1036  | 
hide (open) const mult     (* Multiset.mult, conflicting with monoid.mult *)
  | 
| 
 | 
  1037  | 
*)
  | 
| 
 | 
  1038  | 
  | 
| 
 | 
  1039  | 
subsubsection {* Function definitions *}
 | 
| 
 | 
  1040  | 
  | 
| 
 | 
  1041  | 
constdefs (structure G)
  | 
| 
 | 
  1042  | 
  factors :: "[_, 'a list, 'a] \<Rightarrow> bool"
  | 
| 
 | 
  1043  | 
  "factors G fs a == (\<forall>x \<in> (set fs). irreducible G x) \<and> foldr (op \<otimes>) fs \<one> = a"
  | 
| 
 | 
  1044  | 
  | 
| 
 | 
  1045  | 
  wfactors ::"[_, 'a list, 'a] \<Rightarrow> bool"
  | 
| 
 | 
  1046  | 
  "wfactors G fs a == (\<forall>x \<in> (set fs). irreducible G x) \<and> foldr (op \<otimes>) fs \<one> \<sim> a"
  | 
| 
 | 
  1047  | 
  | 
| 
 | 
  1048  | 
abbreviation
  | 
| 
 | 
  1049  | 
  list_assoc :: "('a,_) monoid_scheme \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "[\<sim>]\<index>" 44) where
 | 
| 
 | 
  1050  | 
  "list_assoc G == list_all2 (op \<sim>\<^bsub>G\<^esub>)"
  | 
| 
 | 
  1051  | 
  | 
| 
 | 
  1052  | 
constdefs (structure G)
  | 
| 
 | 
  1053  | 
  essentially_equal :: "[_, 'a list, 'a list] \<Rightarrow> bool"
  | 
| 
 | 
  1054  | 
  "essentially_equal G fs1 fs2 == (\<exists>fs1'. fs1 <~~> fs1' \<and> fs1' [\<sim>] fs2)"
  | 
| 
 | 
  1055  | 
  | 
| 
 | 
  1056  | 
  | 
| 
 | 
  1057  | 
locale factorial_monoid = comm_monoid_cancel +
  | 
| 
 | 
  1058  | 
  assumes factors_exist: 
  | 
| 
 | 
  1059  | 
          "\<lbrakk>a \<in> carrier G; a \<notin> Units G\<rbrakk> \<Longrightarrow> \<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs a"
  | 
| 
 | 
  1060  | 
      and factors_unique: 
  | 
| 
 | 
  1061  | 
          "\<lbrakk>factors G fs a; factors G fs' a; a \<in> carrier G; a \<notin> Units G; 
  | 
| 
 | 
  1062  | 
            set fs \<subseteq> carrier G; set fs' \<subseteq> carrier G\<rbrakk> \<Longrightarrow> essentially_equal G fs fs'"
  | 
| 
 | 
  1063  | 
  | 
| 
 | 
  1064  | 
  | 
| 
 | 
  1065  | 
subsubsection {* Comparing lists of elements *}
 | 
| 
 | 
  1066  | 
  | 
| 
 | 
  1067  | 
text {* Association on lists *}
 | 
| 
 | 
  1068  | 
  | 
| 
 | 
  1069  | 
lemma (in monoid) listassoc_refl [simp, intro]:
  | 
| 
 | 
  1070  | 
  assumes "set as \<subseteq> carrier G"
  | 
| 
 | 
  1071  | 
  shows "as [\<sim>] as"
  | 
| 
 | 
  1072  | 
using assms
  | 
| 
 | 
  1073  | 
by (induct as) simp+
  | 
| 
 | 
  1074  | 
  | 
| 
 | 
  1075  | 
lemma (in monoid) listassoc_sym [sym]:
  | 
| 
 | 
  1076  | 
  assumes "as [\<sim>] bs"
  | 
| 
 | 
  1077  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1078  | 
  shows "bs [\<sim>] as"
  | 
| 
 | 
  1079  | 
using assms
  | 
| 
 | 
  1080  | 
proof (induct as arbitrary: bs, simp)
  | 
| 
 | 
  1081  | 
  case Cons
  | 
| 
 | 
  1082  | 
  thus ?case
  | 
| 
 | 
  1083  | 
    apply (induct bs, simp)
  | 
| 
 | 
  1084  | 
    apply clarsimp
  | 
| 
 | 
  1085  | 
    apply (iprover intro: associated_sym)
  | 
| 
 | 
  1086  | 
  done
  | 
| 
 | 
  1087  | 
qed
  | 
| 
 | 
  1088  | 
  | 
| 
 | 
  1089  | 
lemma (in monoid) listassoc_trans [trans]:
  | 
| 
 | 
  1090  | 
  assumes "as [\<sim>] bs" and "bs [\<sim>] cs"
  | 
| 
 | 
  1091  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G" and "set cs \<subseteq> carrier G"
  | 
| 
 | 
  1092  | 
  shows "as [\<sim>] cs"
  | 
| 
 | 
  1093  | 
using assms
  | 
| 
 | 
  1094  | 
apply (simp add: list_all2_conv_all_nth set_conv_nth, safe)
  | 
| 
 | 
  1095  | 
apply (rule associated_trans)
  | 
| 
 | 
  1096  | 
    apply (subgoal_tac "as ! i \<sim> bs ! i", assumption)
  | 
| 
 | 
  1097  | 
    apply (simp, simp)
  | 
| 
 | 
  1098  | 
  apply blast+
  | 
| 
 | 
  1099  | 
done
  | 
| 
 | 
  1100  | 
  | 
| 
 | 
  1101  | 
lemma (in monoid_cancel) irrlist_listassoc_cong:
  | 
| 
 | 
  1102  | 
  assumes "\<forall>a\<in>set as. irreducible G a"
  | 
| 
 | 
  1103  | 
    and "as [\<sim>] bs"
  | 
| 
 | 
  1104  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1105  | 
  shows "\<forall>a\<in>set bs. irreducible G a"
  | 
| 
 | 
  1106  | 
using assms
  | 
| 
 | 
  1107  | 
apply (clarsimp simp add: list_all2_conv_all_nth set_conv_nth)
  | 
| 
 | 
  1108  | 
apply (blast intro: irreducible_cong)
  | 
| 
 | 
  1109  | 
done
  | 
| 
 | 
  1110  | 
  | 
| 
 | 
  1111  | 
  | 
| 
 | 
  1112  | 
text {* Permutations *}
 | 
| 
 | 
  1113  | 
  | 
| 
 | 
  1114  | 
lemma perm_map [intro]:
  | 
| 
 | 
  1115  | 
  assumes p: "a <~~> b"
  | 
| 
 | 
  1116  | 
  shows "map f a <~~> map f b"
  | 
| 
 | 
  1117  | 
using p
  | 
| 
 | 
  1118  | 
by induct auto
  | 
| 
 | 
  1119  | 
  | 
| 
 | 
  1120  | 
lemma perm_map_switch:
  | 
| 
 | 
  1121  | 
  assumes m: "map f a = map f b" and p: "b <~~> c"
  | 
| 
 | 
  1122  | 
  shows "\<exists>d. a <~~> d \<and> map f d = map f c"
  | 
| 
 | 
  1123  | 
using p m
  | 
| 
 | 
  1124  | 
by (induct arbitrary: a) (simp, force, force, blast)
  | 
| 
 | 
  1125  | 
  | 
| 
 | 
  1126  | 
lemma (in monoid) perm_assoc_switch:
  | 
| 
 | 
  1127  | 
   assumes a:"as [\<sim>] bs" and p: "bs <~~> cs"
  | 
| 
 | 
  1128  | 
   shows "\<exists>bs'. as <~~> bs' \<and> bs' [\<sim>] cs"
  | 
| 
 | 
  1129  | 
using p a
  | 
| 
 | 
  1130  | 
apply (induct bs cs arbitrary: as, simp)
  | 
| 
 | 
  1131  | 
  apply (clarsimp simp add: list_all2_Cons2, blast)
  | 
| 
 | 
  1132  | 
 apply (clarsimp simp add: list_all2_Cons2)
  | 
| 
 | 
  1133  | 
 apply blast
  | 
| 
 | 
  1134  | 
apply blast
  | 
| 
 | 
  1135  | 
done
  | 
| 
 | 
  1136  | 
  | 
| 
 | 
  1137  | 
lemma (in monoid) perm_assoc_switch_r:
  | 
| 
 | 
  1138  | 
   assumes p: "as <~~> bs" and a:"bs [\<sim>] cs"
  | 
| 
 | 
  1139  | 
   shows "\<exists>bs'. as [\<sim>] bs' \<and> bs' <~~> cs"
  | 
| 
 | 
  1140  | 
using p a
  | 
| 
 | 
  1141  | 
apply (induct as bs arbitrary: cs, simp)
  | 
| 
 | 
  1142  | 
  apply (clarsimp simp add: list_all2_Cons1, blast)
  | 
| 
 | 
  1143  | 
 apply (clarsimp simp add: list_all2_Cons1)
  | 
| 
 | 
  1144  | 
 apply blast
  | 
| 
 | 
  1145  | 
apply blast
  | 
| 
 | 
  1146  | 
done
  | 
| 
 | 
  1147  | 
  | 
| 
 | 
  1148  | 
declare perm_sym [sym]
  | 
| 
 | 
  1149  | 
  | 
| 
 | 
  1150  | 
lemma perm_setP:
  | 
| 
 | 
  1151  | 
  assumes perm: "as <~~> bs"
  | 
| 
 | 
  1152  | 
    and as: "P (set as)"
  | 
| 
 | 
  1153  | 
  shows "P (set bs)"
  | 
| 
 | 
  1154  | 
proof -
  | 
| 
 | 
  1155  | 
  from perm
  | 
| 
 | 
  1156  | 
      have "multiset_of as = multiset_of bs"
  | 
| 
 | 
  1157  | 
      by (simp add: multiset_of_eq_perm)
  | 
| 
 | 
  1158  | 
  hence "set as = set bs" by (rule multiset_of_eq_setD)
  | 
| 
 | 
  1159  | 
  with as
  | 
| 
 | 
  1160  | 
      show "P (set bs)" by simp
  | 
| 
 | 
  1161  | 
qed
  | 
| 
 | 
  1162  | 
  | 
| 
 | 
  1163  | 
lemmas (in monoid) perm_closed =
  | 
| 
 | 
  1164  | 
    perm_setP[of _ _ "\<lambda>as. as \<subseteq> carrier G"]
  | 
| 
 | 
  1165  | 
  | 
| 
 | 
  1166  | 
lemmas (in monoid) irrlist_perm_cong =
  | 
| 
 | 
  1167  | 
    perm_setP[of _ _ "\<lambda>as. \<forall>a\<in>as. irreducible G a"]
  | 
| 
 | 
  1168  | 
  | 
| 
 | 
  1169  | 
  | 
| 
 | 
  1170  | 
text {* Essentially equal factorizations *}
 | 
| 
 | 
  1171  | 
  | 
| 
 | 
  1172  | 
lemma (in monoid) essentially_equalI:
  | 
| 
 | 
  1173  | 
  assumes ex: "fs1 <~~> fs1'"  "fs1' [\<sim>] fs2"
  | 
| 
 | 
  1174  | 
  shows "essentially_equal G fs1 fs2"
  | 
| 
 | 
  1175  | 
using ex
  | 
| 
 | 
  1176  | 
unfolding essentially_equal_def
  | 
| 
 | 
  1177  | 
by fast
  | 
| 
 | 
  1178  | 
  | 
| 
 | 
  1179  | 
lemma (in monoid) essentially_equalE:
  | 
| 
 | 
  1180  | 
  assumes ee: "essentially_equal G fs1 fs2"
  | 
| 
 | 
  1181  | 
    and e: "\<And>fs1'. \<lbrakk>fs1 <~~> fs1'; fs1' [\<sim>] fs2\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
  1182  | 
  shows "P"
  | 
| 
 | 
  1183  | 
using ee
  | 
| 
 | 
  1184  | 
unfolding essentially_equal_def
  | 
| 
 | 
  1185  | 
by (fast intro: e)
  | 
| 
 | 
  1186  | 
  | 
| 
 | 
  1187  | 
lemma (in monoid) ee_refl [simp,intro]:
  | 
| 
 | 
  1188  | 
  assumes carr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1189  | 
  shows "essentially_equal G as as"
  | 
| 
 | 
  1190  | 
using carr
  | 
| 
 | 
  1191  | 
by (fast intro: essentially_equalI)
  | 
| 
 | 
  1192  | 
  | 
| 
 | 
  1193  | 
lemma (in monoid) ee_sym [sym]:
  | 
| 
 | 
  1194  | 
  assumes ee: "essentially_equal G as bs"
  | 
| 
 | 
  1195  | 
    and carr: "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1196  | 
  shows "essentially_equal G bs as"
  | 
| 
 | 
  1197  | 
using ee
  | 
| 
 | 
  1198  | 
proof (elim essentially_equalE)
  | 
| 
 | 
  1199  | 
  fix fs
  | 
| 
 | 
  1200  | 
  assume "as <~~> fs"  "fs [\<sim>] bs"
  | 
| 
 | 
  1201  | 
  hence "\<exists>fs'. as [\<sim>] fs' \<and> fs' <~~> bs" by (rule perm_assoc_switch_r)
  | 
| 
 | 
  1202  | 
  from this obtain fs'
  | 
| 
 | 
  1203  | 
      where a: "as [\<sim>] fs'" and p: "fs' <~~> bs"
  | 
| 
 | 
  1204  | 
      by auto
  | 
| 
 | 
  1205  | 
  from p have "bs <~~> fs'" by (rule perm_sym)
  | 
| 
 | 
  1206  | 
  with a[symmetric] carr
  | 
| 
 | 
  1207  | 
      show ?thesis
  | 
| 
 | 
  1208  | 
      by (iprover intro: essentially_equalI perm_closed)
  | 
| 
 | 
  1209  | 
qed
  | 
| 
 | 
  1210  | 
  | 
| 
 | 
  1211  | 
lemma (in monoid) ee_trans [trans]:
  | 
| 
 | 
  1212  | 
  assumes ab: "essentially_equal G as bs" and bc: "essentially_equal G bs cs"
  | 
| 
 | 
  1213  | 
    and ascarr: "set as \<subseteq> carrier G" 
  | 
| 
 | 
  1214  | 
    and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1215  | 
    and cscarr: "set cs \<subseteq> carrier G"
  | 
| 
 | 
  1216  | 
  shows "essentially_equal G as cs"
  | 
| 
 | 
  1217  | 
using ab bc
  | 
| 
 | 
  1218  | 
proof (elim essentially_equalE)
  | 
| 
 | 
  1219  | 
  fix abs bcs
  | 
| 
 | 
  1220  | 
  assume  "abs [\<sim>] bs" and pb: "bs <~~> bcs"
  | 
| 
 | 
  1221  | 
  hence "\<exists>bs'. abs <~~> bs' \<and> bs' [\<sim>] bcs" by (rule perm_assoc_switch)
  | 
| 
 | 
  1222  | 
  from this obtain bs'
  | 
| 
 | 
  1223  | 
      where p: "abs <~~> bs'" and a: "bs' [\<sim>] bcs"
  | 
| 
 | 
  1224  | 
      by auto
  | 
| 
 | 
  1225  | 
  | 
| 
 | 
  1226  | 
  assume "as <~~> abs"
  | 
| 
 | 
  1227  | 
  with p
  | 
| 
 | 
  1228  | 
      have pp: "as <~~> bs'" by fast
  | 
| 
 | 
  1229  | 
  | 
| 
 | 
  1230  | 
  from pp ascarr have c1: "set bs' \<subseteq> carrier G" by (rule perm_closed)
  | 
| 
 | 
  1231  | 
  from pb bscarr have c2: "set bcs \<subseteq> carrier G" by (rule perm_closed)
  | 
| 
 | 
  1232  | 
  note a
  | 
| 
 | 
  1233  | 
  also assume "bcs [\<sim>] cs"
  | 
| 
 | 
  1234  | 
  finally (listassoc_trans) have"bs' [\<sim>] cs" by (simp add: c1 c2 cscarr)
  | 
| 
 | 
  1235  | 
  | 
| 
 | 
  1236  | 
  with pp
  | 
| 
 | 
  1237  | 
      show ?thesis
  | 
| 
 | 
  1238  | 
      by (rule essentially_equalI)
  | 
| 
 | 
  1239  | 
qed
  | 
| 
 | 
  1240  | 
  | 
| 
 | 
  1241  | 
  | 
| 
 | 
  1242  | 
subsubsection {* Properties of lists of elements *}
 | 
| 
 | 
  1243  | 
  | 
| 
 | 
  1244  | 
text {* Multiplication of factors in a list *}
 | 
| 
 | 
  1245  | 
  | 
| 
 | 
  1246  | 
lemma (in monoid) multlist_closed [simp, intro]:
  | 
| 
 | 
  1247  | 
  assumes ascarr: "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1248  | 
  shows "foldr (op \<otimes>) fs \<one> \<in> carrier G"
  | 
| 
 | 
  1249  | 
by (insert ascarr, induct fs, simp+)
  | 
| 
 | 
  1250  | 
  | 
| 
 | 
  1251  | 
lemma  (in comm_monoid) multlist_dividesI (*[intro]*):
  | 
| 
 | 
  1252  | 
  assumes "f \<in> set fs" and "f \<in> carrier G" and "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1253  | 
  shows "f divides (foldr (op \<otimes>) fs \<one>)"
  | 
| 
 | 
  1254  | 
using assms
  | 
| 
 | 
  1255  | 
apply (induct fs)
  | 
| 
 | 
  1256  | 
 apply simp
  | 
| 
 | 
  1257  | 
apply (case_tac "f = a", simp)
  | 
| 
 | 
  1258  | 
 apply (fast intro: dividesI)
  | 
| 
 | 
  1259  | 
apply clarsimp
  | 
| 
 | 
  1260  | 
apply (elim dividesE, intro dividesI)
  | 
| 
 | 
  1261  | 
 defer 1
  | 
| 
 | 
  1262  | 
 apply (simp add: m_comm)
  | 
| 
 | 
  1263  | 
 apply (simp add: m_assoc[symmetric])
  | 
| 
 | 
  1264  | 
 apply (simp add: m_comm)
  | 
| 
 | 
  1265  | 
apply simp
  | 
| 
 | 
  1266  | 
done
  | 
| 
 | 
  1267  | 
  | 
| 
 | 
  1268  | 
lemma (in comm_monoid_cancel) multlist_listassoc_cong:
  | 
| 
 | 
  1269  | 
  assumes "fs [\<sim>] fs'"
  | 
| 
 | 
  1270  | 
    and "set fs \<subseteq> carrier G" and "set fs' \<subseteq> carrier G"
  | 
| 
 | 
  1271  | 
  shows "foldr (op \<otimes>) fs \<one> \<sim> foldr (op \<otimes>) fs' \<one>"
  | 
| 
 | 
  1272  | 
using assms
  | 
| 
 | 
  1273  | 
proof (induct fs arbitrary: fs', simp)
  | 
| 
 | 
  1274  | 
  case (Cons a as fs')
  | 
| 
 | 
  1275  | 
  thus ?case
  | 
| 
 | 
  1276  | 
  apply (induct fs', simp)
  | 
| 
 | 
  1277  | 
  proof clarsimp
  | 
| 
 | 
  1278  | 
    fix b bs
  | 
| 
 | 
  1279  | 
    assume "a \<sim> b" 
  | 
| 
 | 
  1280  | 
      and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  1281  | 
      and ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1282  | 
    hence p: "a \<otimes> foldr op \<otimes> as \<one> \<sim> b \<otimes> foldr op \<otimes> as \<one>"
  | 
| 
 | 
  1283  | 
        by (fast intro: mult_cong_l)
  | 
| 
 | 
  1284  | 
    also
  | 
| 
 | 
  1285  | 
      assume "as [\<sim>] bs"
  | 
| 
 | 
  1286  | 
         and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1287  | 
         and "\<And>fs'. \<lbrakk>as [\<sim>] fs'; set fs' \<subseteq> carrier G\<rbrakk> \<Longrightarrow> foldr op \<otimes> as \<one> \<sim> foldr op \<otimes> fs' \<one>"
  | 
| 
 | 
  1288  | 
      hence "foldr op \<otimes> as \<one> \<sim> foldr op \<otimes> bs \<one>" by simp
  | 
| 
 | 
  1289  | 
      with ascarr bscarr bcarr
  | 
| 
 | 
  1290  | 
          have "b \<otimes> foldr op \<otimes> as \<one> \<sim> b \<otimes> foldr op \<otimes> bs \<one>"
  | 
| 
 | 
  1291  | 
          by (fast intro: mult_cong_r)
  | 
| 
 | 
  1292  | 
   finally
  | 
| 
 | 
  1293  | 
       show "a \<otimes> foldr op \<otimes> as \<one> \<sim> b \<otimes> foldr op \<otimes> bs \<one>"
  | 
| 
 | 
  1294  | 
       by (simp add: ascarr bscarr acarr bcarr)
  | 
| 
 | 
  1295  | 
  qed
  | 
| 
 | 
  1296  | 
qed
  | 
| 
 | 
  1297  | 
  | 
| 
 | 
  1298  | 
lemma (in comm_monoid) multlist_perm_cong:
  | 
| 
 | 
  1299  | 
  assumes prm: "as <~~> bs"
  | 
| 
 | 
  1300  | 
    and ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1301  | 
  shows "foldr (op \<otimes>) as \<one> = foldr (op \<otimes>) bs \<one>"
  | 
| 
 | 
  1302  | 
using prm ascarr
  | 
| 
 | 
  1303  | 
apply (induct, simp, clarsimp simp add: m_ac, clarsimp)
  | 
| 
 | 
  1304  | 
proof clarsimp
  | 
| 
 | 
  1305  | 
  fix xs ys zs
  | 
| 
 | 
  1306  | 
  assume "xs <~~> ys"  "set xs \<subseteq> carrier G"
  | 
| 
 | 
  1307  | 
  hence "set ys \<subseteq> carrier G" by (rule perm_closed)
  | 
| 
 | 
  1308  | 
  moreover assume "set ys \<subseteq> carrier G \<Longrightarrow> foldr op \<otimes> ys \<one> = foldr op \<otimes> zs \<one>"
  | 
| 
 | 
  1309  | 
  ultimately show "foldr op \<otimes> ys \<one> = foldr op \<otimes> zs \<one>" by simp
  | 
| 
 | 
  1310  | 
qed
  | 
| 
 | 
  1311  | 
  | 
| 
 | 
  1312  | 
lemma (in comm_monoid_cancel) multlist_ee_cong:
  | 
| 
 | 
  1313  | 
  assumes "essentially_equal G fs fs'"
  | 
| 
 | 
  1314  | 
    and "set fs \<subseteq> carrier G" and "set fs' \<subseteq> carrier G"
  | 
| 
 | 
  1315  | 
  shows "foldr (op \<otimes>) fs \<one> \<sim> foldr (op \<otimes>) fs' \<one>"
  | 
| 
 | 
  1316  | 
using assms
  | 
| 
 | 
  1317  | 
apply (elim essentially_equalE)
  | 
| 
 | 
  1318  | 
apply (simp add: multlist_perm_cong multlist_listassoc_cong perm_closed)
  | 
| 
 | 
  1319  | 
done
  | 
| 
 | 
  1320  | 
  | 
| 
 | 
  1321  | 
  | 
| 
 | 
  1322  | 
subsubsection {* Factorization in irreducible elements *}
 | 
| 
 | 
  1323  | 
  | 
| 
 | 
  1324  | 
lemma wfactorsI:
  | 
| 
 | 
  1325  | 
  includes (struct G)
  | 
| 
 | 
  1326  | 
  assumes "\<forall>f\<in>set fs. irreducible G f"
  | 
| 
 | 
  1327  | 
    and "foldr (op \<otimes>) fs \<one> \<sim> a"
  | 
| 
 | 
  1328  | 
  shows "wfactors G fs a"
  | 
| 
 | 
  1329  | 
using assms
  | 
| 
 | 
  1330  | 
unfolding wfactors_def
  | 
| 
 | 
  1331  | 
by simp
  | 
| 
 | 
  1332  | 
  | 
| 
 | 
  1333  | 
lemma wfactorsE:
  | 
| 
 | 
  1334  | 
  includes (struct G)
  | 
| 
 | 
  1335  | 
  assumes wf: "wfactors G fs a"
  | 
| 
 | 
  1336  | 
    and e: "\<lbrakk>\<forall>f\<in>set fs. irreducible G f; foldr (op \<otimes>) fs \<one> \<sim> a\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
  1337  | 
  shows "P"
  | 
| 
 | 
  1338  | 
using wf
  | 
| 
 | 
  1339  | 
unfolding wfactors_def
  | 
| 
 | 
  1340  | 
by (fast dest: e)
  | 
| 
 | 
  1341  | 
  | 
| 
 | 
  1342  | 
lemma (in monoid) factorsI:
  | 
| 
 | 
  1343  | 
  includes (struct G)
  | 
| 
 | 
  1344  | 
  assumes "\<forall>f\<in>set fs. irreducible G f"
  | 
| 
 | 
  1345  | 
    and "foldr (op \<otimes>) fs \<one> = a"
  | 
| 
 | 
  1346  | 
  shows "factors G fs a"
  | 
| 
 | 
  1347  | 
using assms
  | 
| 
 | 
  1348  | 
unfolding factors_def
  | 
| 
 | 
  1349  | 
by simp
  | 
| 
 | 
  1350  | 
  | 
| 
 | 
  1351  | 
lemma factorsE:
  | 
| 
 | 
  1352  | 
  includes (struct G)
  | 
| 
 | 
  1353  | 
  assumes f: "factors G fs a"
  | 
| 
 | 
  1354  | 
    and e: "\<lbrakk>\<forall>f\<in>set fs. irreducible G f; foldr (op \<otimes>) fs \<one> = a\<rbrakk> \<Longrightarrow> P"
  | 
| 
 | 
  1355  | 
  shows "P"
  | 
| 
 | 
  1356  | 
using f
  | 
| 
 | 
  1357  | 
unfolding factors_def
  | 
| 
 | 
  1358  | 
by (simp add: e)
  | 
| 
 | 
  1359  | 
  | 
| 
 | 
  1360  | 
lemma (in monoid) factors_wfactors:
  | 
| 
 | 
  1361  | 
  assumes "factors G as a" and "set as \<subseteq> carrier G"
  | 
| 
 | 
  1362  | 
  shows "wfactors G as a"
  | 
| 
 | 
  1363  | 
using assms
  | 
| 
 | 
  1364  | 
by (blast elim: factorsE intro: wfactorsI)
  | 
| 
 | 
  1365  | 
  | 
| 
 | 
  1366  | 
lemma (in monoid) wfactors_factors:
  | 
| 
 | 
  1367  | 
  assumes "wfactors G as a" and "set as \<subseteq> carrier G"
  | 
| 
 | 
  1368  | 
  shows "\<exists>a'. factors G as a' \<and> a' \<sim> a"
  | 
| 
 | 
  1369  | 
using assms
  | 
| 
 | 
  1370  | 
by (blast elim: wfactorsE intro: factorsI)
  | 
| 
 | 
  1371  | 
  | 
| 
 | 
  1372  | 
lemma (in monoid) factors_closed [dest]:
  | 
| 
 | 
  1373  | 
  assumes "factors G fs a" and "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1374  | 
  shows "a \<in> carrier G"
  | 
| 
 | 
  1375  | 
using assms
  | 
| 
 | 
  1376  | 
by (elim factorsE, clarsimp)
  | 
| 
 | 
  1377  | 
  | 
| 
 | 
  1378  | 
lemma (in monoid) nunit_factors:
  | 
| 
 | 
  1379  | 
  assumes anunit: "a \<notin> Units G"
  | 
| 
 | 
  1380  | 
    and fs: "factors G as a"
  | 
| 
 | 
  1381  | 
  shows "length as > 0"
  | 
| 
 | 
  1382  | 
apply (insert fs, elim factorsE)
  | 
| 
 | 
  1383  | 
proof (cases "length as = 0")
  | 
| 
 | 
  1384  | 
  assume "length as = 0"
  | 
| 
 | 
  1385  | 
  hence fold: "foldr op \<otimes> as \<one> = \<one>" by force
  | 
| 
 | 
  1386  | 
  | 
| 
 | 
  1387  | 
  assume "foldr op \<otimes> as \<one> = a"
  | 
| 
 | 
  1388  | 
  with fold
  | 
| 
 | 
  1389  | 
       have "a = \<one>" by simp
  | 
| 
 | 
  1390  | 
  then have "a \<in> Units G" by fast
  | 
| 
 | 
  1391  | 
  with anunit
  | 
| 
 | 
  1392  | 
       have "False" by simp
  | 
| 
 | 
  1393  | 
  thus ?thesis ..
  | 
| 
 | 
  1394  | 
qed simp
  | 
| 
 | 
  1395  | 
  | 
| 
 | 
  1396  | 
lemma (in monoid) unit_wfactors [simp]:
  | 
| 
 | 
  1397  | 
  assumes aunit: "a \<in> Units G"
  | 
| 
 | 
  1398  | 
  shows "wfactors G [] a"
  | 
| 
 | 
  1399  | 
using aunit
  | 
| 
 | 
  1400  | 
by (intro wfactorsI) (simp, simp add: Units_assoc)
  | 
| 
 | 
  1401  | 
  | 
| 
 | 
  1402  | 
lemma (in comm_monoid_cancel) unit_wfactors_empty:
  | 
| 
 | 
  1403  | 
  assumes aunit: "a \<in> Units G"
  | 
| 
 | 
  1404  | 
    and wf: "wfactors G fs a"
  | 
| 
 | 
  1405  | 
    and carr[simp]: "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1406  | 
  shows "fs = []"
  | 
| 
 | 
  1407  | 
proof (rule ccontr, cases fs, simp)
  | 
| 
 | 
  1408  | 
  fix f fs'
  | 
| 
 | 
  1409  | 
  assume fs: "fs = f # fs'"
  | 
| 
 | 
  1410  | 
  | 
| 
 | 
  1411  | 
  from carr
  | 
| 
 | 
  1412  | 
      have fcarr[simp]: "f \<in> carrier G"
  | 
| 
 | 
  1413  | 
      and carr'[simp]: "set fs' \<subseteq> carrier G"
  | 
| 
 | 
  1414  | 
      by (simp add: fs)+
  | 
| 
 | 
  1415  | 
  | 
| 
 | 
  1416  | 
  from fs wf
  | 
| 
 | 
  1417  | 
      have "irreducible G f" by (simp add: wfactors_def)
  | 
| 
 | 
  1418  | 
  hence fnunit: "f \<notin> Units G" by (fast elim: irreducibleE)
  | 
| 
 | 
  1419  | 
  | 
| 
 | 
  1420  | 
  from fs wf
  | 
| 
 | 
  1421  | 
      have a: "f \<otimes> foldr (op \<otimes>) fs' \<one> \<sim> a" by (simp add: wfactors_def)
  | 
| 
 | 
  1422  | 
  | 
| 
 | 
  1423  | 
  note aunit
  | 
| 
 | 
  1424  | 
  also from fs wf
  | 
| 
 | 
  1425  | 
       have a: "f \<otimes> foldr (op \<otimes>) fs' \<one> \<sim> a" by (simp add: wfactors_def)
  | 
| 
 | 
  1426  | 
       have "a \<sim> f \<otimes> foldr (op \<otimes>) fs' \<one>" 
  | 
| 
 | 
  1427  | 
       by (simp add: Units_closed[OF aunit] a[symmetric])
  | 
| 
 | 
  1428  | 
  finally
  | 
| 
 | 
  1429  | 
       have "f \<otimes> foldr (op \<otimes>) fs' \<one> \<in> Units G" by simp
  | 
| 
 | 
  1430  | 
  hence "f \<in> Units G" by (intro unit_factor[of f], simp+)
  | 
| 
 | 
  1431  | 
  | 
| 
 | 
  1432  | 
  with fnunit show "False" by simp
  | 
| 
 | 
  1433  | 
qed
  | 
| 
 | 
  1434  | 
  | 
| 
 | 
  1435  | 
  | 
| 
 | 
  1436  | 
text {* Comparing wfactors *}
 | 
| 
 | 
  1437  | 
  | 
| 
 | 
  1438  | 
lemma (in comm_monoid_cancel) wfactors_listassoc_cong_l:
  | 
| 
 | 
  1439  | 
  assumes fact: "wfactors G fs a"
  | 
| 
 | 
  1440  | 
    and asc: "fs [\<sim>] fs'"
  | 
| 
 | 
  1441  | 
    and carr: "a \<in> carrier G"  "set fs \<subseteq> carrier G"  "set fs' \<subseteq> carrier G"
  | 
| 
 | 
  1442  | 
  shows "wfactors G fs' a"
  | 
| 
 | 
  1443  | 
using fact
  | 
| 
 | 
  1444  | 
apply (elim wfactorsE, intro wfactorsI)
  | 
| 
 | 
  1445  | 
proof -
  | 
| 
 | 
  1446  | 
  assume "\<forall>f\<in>set fs. irreducible G f"
  | 
| 
 | 
  1447  | 
  also note asc
  | 
| 
 | 
  1448  | 
  finally (irrlist_listassoc_cong)
  | 
| 
 | 
  1449  | 
       show "\<forall>f\<in>set fs'. irreducible G f" by (simp add: carr)
  | 
| 
 | 
  1450  | 
next
  | 
| 
 | 
  1451  | 
  from asc[symmetric]
  | 
| 
 | 
  1452  | 
       have "foldr op \<otimes> fs' \<one> \<sim> foldr op \<otimes> fs \<one>" 
  | 
| 
 | 
  1453  | 
       by (simp add: multlist_listassoc_cong carr)
  | 
| 
 | 
  1454  | 
  also assume "foldr op \<otimes> fs \<one> \<sim> a"
  | 
| 
 | 
  1455  | 
  finally
  | 
| 
 | 
  1456  | 
       show "foldr op \<otimes> fs' \<one> \<sim> a" by (simp add: carr)
  | 
| 
 | 
  1457  | 
qed
  | 
| 
 | 
  1458  | 
  | 
| 
 | 
  1459  | 
lemma (in comm_monoid) wfactors_perm_cong_l:
  | 
| 
 | 
  1460  | 
  assumes "wfactors G fs a"
  | 
| 
 | 
  1461  | 
    and "fs <~~> fs'"
  | 
| 
 | 
  1462  | 
    and "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1463  | 
  shows "wfactors G fs' a"
  | 
| 
 | 
  1464  | 
using assms
  | 
| 
 | 
  1465  | 
apply (elim wfactorsE, intro wfactorsI)
  | 
| 
 | 
  1466  | 
 apply (rule irrlist_perm_cong, assumption+)
  | 
| 
 | 
  1467  | 
apply (simp add: multlist_perm_cong[symmetric])
  | 
| 
 | 
  1468  | 
done
  | 
| 
 | 
  1469  | 
  | 
| 
 | 
  1470  | 
lemma (in comm_monoid_cancel) wfactors_ee_cong_l [trans]:
  | 
| 
 | 
  1471  | 
  assumes ee: "essentially_equal G as bs"
  | 
| 
 | 
  1472  | 
    and bfs: "wfactors G bs b"
  | 
| 
 | 
  1473  | 
    and carr: "b \<in> carrier G"  "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1474  | 
  shows "wfactors G as b"
  | 
| 
 | 
  1475  | 
using ee
  | 
| 
 | 
  1476  | 
proof (elim essentially_equalE)
  | 
| 
 | 
  1477  | 
  fix fs
  | 
| 
 | 
  1478  | 
  assume prm: "as <~~> fs"
  | 
| 
 | 
  1479  | 
  with carr
  | 
| 
 | 
  1480  | 
       have fscarr: "set fs \<subseteq> carrier G" by (simp add: perm_closed)
  | 
| 
 | 
  1481  | 
  | 
| 
 | 
  1482  | 
  note bfs
  | 
| 
 | 
  1483  | 
  also assume [symmetric]: "fs [\<sim>] bs"
  | 
| 
 | 
  1484  | 
  also (wfactors_listassoc_cong_l)
  | 
| 
 | 
  1485  | 
       note prm[symmetric]
  | 
| 
 | 
  1486  | 
  finally (wfactors_perm_cong_l)
  | 
| 
 | 
  1487  | 
       show "wfactors G as b" by (simp add: carr fscarr)
  | 
| 
 | 
  1488  | 
qed
  | 
| 
 | 
  1489  | 
  | 
| 
 | 
  1490  | 
lemma (in monoid) wfactors_cong_r [trans]:
  | 
| 
 | 
  1491  | 
  assumes fac: "wfactors G fs a" and aa': "a \<sim> a'"
  | 
| 
 | 
  1492  | 
    and carr[simp]: "a \<in> carrier G"  "a' \<in> carrier G"  "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1493  | 
  shows "wfactors G fs a'"
  | 
| 
 | 
  1494  | 
using fac
  | 
| 
 | 
  1495  | 
proof (elim wfactorsE, intro wfactorsI)
  | 
| 
 | 
  1496  | 
  assume "foldr op \<otimes> fs \<one> \<sim> a" also note aa'
  | 
| 
 | 
  1497  | 
  finally show "foldr op \<otimes> fs \<one> \<sim> a'" by simp
  | 
| 
 | 
  1498  | 
qed
  | 
| 
 | 
  1499  | 
  | 
| 
 | 
  1500  | 
  | 
| 
 | 
  1501  | 
subsubsection {* Essentially equal factorizations *}
 | 
| 
 | 
  1502  | 
  | 
| 
 | 
  1503  | 
lemma (in comm_monoid_cancel) unitfactor_ee:
  | 
| 
 | 
  1504  | 
  assumes uunit: "u \<in> Units G"
  | 
| 
 | 
  1505  | 
    and carr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1506  | 
  shows "essentially_equal G (as[0 := (as!0 \<otimes> u)]) as" (is "essentially_equal G ?as' as")
  | 
| 
 | 
  1507  | 
using assms
  | 
| 
 | 
  1508  | 
apply (intro essentially_equalI[of _ ?as'], simp)
  | 
| 
 | 
  1509  | 
apply (cases as, simp)
  | 
| 
 | 
  1510  | 
apply (clarsimp, fast intro: associatedI2[of u])
  | 
| 
 | 
  1511  | 
done
  | 
| 
 | 
  1512  | 
  | 
| 
 | 
  1513  | 
lemma (in comm_monoid_cancel) factors_cong_unit:
  | 
| 
 | 
  1514  | 
  assumes uunit: "u \<in> Units G" and anunit: "a \<notin> Units G"
  | 
| 
 | 
  1515  | 
    and afs: "factors G as a"
  | 
| 
 | 
  1516  | 
    and ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1517  | 
  shows "factors G (as[0 := (as!0 \<otimes> u)]) (a \<otimes> u)" (is "factors G ?as' ?a'")
  | 
| 
 | 
  1518  | 
using assms
  | 
| 
 | 
  1519  | 
apply (elim factorsE, clarify)
  | 
| 
 | 
  1520  | 
apply (cases as)
  | 
| 
 | 
  1521  | 
 apply (simp add: nunit_factors)
  | 
| 
 | 
  1522  | 
apply clarsimp
  | 
| 
 | 
  1523  | 
apply (elim factorsE, intro factorsI)
  | 
| 
 | 
  1524  | 
 apply (clarsimp, fast intro: irreducible_prod_rI)
  | 
| 
 | 
  1525  | 
apply (simp add: m_ac Units_closed)
  | 
| 
 | 
  1526  | 
done
  | 
| 
 | 
  1527  | 
  | 
| 
 | 
  1528  | 
lemma (in comm_monoid) perm_wfactorsD:
  | 
| 
 | 
  1529  | 
  assumes prm: "as <~~> bs"
  | 
| 
 | 
  1530  | 
    and afs: "wfactors G as a" and bfs: "wfactors G bs b"
  | 
| 
 | 
  1531  | 
    and [simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  1532  | 
    and ascarr[simp]: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1533  | 
  shows "a \<sim> b"
  | 
| 
 | 
  1534  | 
using afs bfs
  | 
| 
 | 
  1535  | 
proof (elim wfactorsE)
  | 
| 
 | 
  1536  | 
  from prm have [simp]: "set bs \<subseteq> carrier G" by (simp add: perm_closed)
  | 
| 
 | 
  1537  | 
  assume "foldr op \<otimes> as \<one> \<sim> a"
  | 
| 
 | 
  1538  | 
  hence "a \<sim> foldr op \<otimes> as \<one>" by (rule associated_sym, simp+)
  | 
| 
 | 
  1539  | 
  also from prm
  | 
| 
 | 
  1540  | 
       have "foldr op \<otimes> as \<one> = foldr op \<otimes> bs \<one>" by (rule multlist_perm_cong, simp)
  | 
| 
 | 
  1541  | 
  also assume "foldr op \<otimes> bs \<one> \<sim> b"
  | 
| 
 | 
  1542  | 
  finally
  | 
| 
 | 
  1543  | 
       show "a \<sim> b" by simp
  | 
| 
 | 
  1544  | 
qed
  | 
| 
 | 
  1545  | 
  | 
| 
 | 
  1546  | 
lemma (in comm_monoid_cancel) listassoc_wfactorsD:
  | 
| 
 | 
  1547  | 
  assumes assoc: "as [\<sim>] bs"
  | 
| 
 | 
  1548  | 
    and afs: "wfactors G as a" and bfs: "wfactors G bs b"
  | 
| 
 | 
  1549  | 
    and [simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  1550  | 
    and [simp]: "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1551  | 
  shows "a \<sim> b"
  | 
| 
 | 
  1552  | 
using afs bfs
  | 
| 
 | 
  1553  | 
proof (elim wfactorsE)
  | 
| 
 | 
  1554  | 
  assume "foldr op \<otimes> as \<one> \<sim> a"
  | 
| 
 | 
  1555  | 
  hence "a \<sim> foldr op \<otimes> as \<one>" by (rule associated_sym, simp+)
  | 
| 
 | 
  1556  | 
  also from assoc
  | 
| 
 | 
  1557  | 
       have "foldr op \<otimes> as \<one> \<sim> foldr op \<otimes> bs \<one>" by (rule multlist_listassoc_cong, simp+)
  | 
| 
 | 
  1558  | 
  also assume "foldr op \<otimes> bs \<one> \<sim> b"
  | 
| 
 | 
  1559  | 
  finally
  | 
| 
 | 
  1560  | 
       show "a \<sim> b" by simp
  | 
| 
 | 
  1561  | 
qed
  | 
| 
 | 
  1562  | 
  | 
| 
 | 
  1563  | 
lemma (in comm_monoid_cancel) ee_wfactorsD:
  | 
| 
 | 
  1564  | 
  assumes ee: "essentially_equal G as bs"
  | 
| 
 | 
  1565  | 
    and afs: "wfactors G as a" and bfs: "wfactors G bs b"
  | 
| 
 | 
  1566  | 
    and [simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  1567  | 
    and ascarr[simp]: "set as \<subseteq> carrier G" and bscarr[simp]: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1568  | 
  shows "a \<sim> b"
  | 
| 
 | 
  1569  | 
using ee
  | 
| 
 | 
  1570  | 
proof (elim essentially_equalE)
  | 
| 
 | 
  1571  | 
  fix fs
  | 
| 
 | 
  1572  | 
  assume prm: "as <~~> fs"
  | 
| 
 | 
  1573  | 
  hence as'carr[simp]: "set fs \<subseteq> carrier G" by (simp add: perm_closed)
  | 
| 
 | 
  1574  | 
  from afs prm
  | 
| 
 | 
  1575  | 
      have afs': "wfactors G fs a" by (rule wfactors_perm_cong_l, simp)
  | 
| 
 | 
  1576  | 
  assume "fs [\<sim>] bs"
  | 
| 
 | 
  1577  | 
  from this afs' bfs
  | 
| 
 | 
  1578  | 
      show "a \<sim> b" by (rule listassoc_wfactorsD, simp+)
  | 
| 
 | 
  1579  | 
qed
  | 
| 
 | 
  1580  | 
  | 
| 
 | 
  1581  | 
lemma (in comm_monoid_cancel) ee_factorsD:
  | 
| 
 | 
  1582  | 
  assumes ee: "essentially_equal G as bs"
  | 
| 
 | 
  1583  | 
    and afs: "factors G as a" and bfs:"factors G bs b"
  | 
| 
 | 
  1584  | 
    and "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1585  | 
  shows "a \<sim> b"
  | 
| 
 | 
  1586  | 
using assms
  | 
| 
 | 
  1587  | 
by (blast intro: factors_wfactors dest: ee_wfactorsD)
  | 
| 
 | 
  1588  | 
  | 
| 
 | 
  1589  | 
lemma (in factorial_monoid) ee_factorsI:
  | 
| 
 | 
  1590  | 
  assumes ab: "a \<sim> b"
  | 
| 
 | 
  1591  | 
    and afs: "factors G as a" and anunit: "a \<notin> Units G"
  | 
| 
 | 
  1592  | 
    and bfs: "factors G bs b" and bnunit: "b \<notin> Units G"
  | 
| 
 | 
  1593  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1594  | 
  shows "essentially_equal G as bs"
  | 
| 
 | 
  1595  | 
proof -
  | 
| 
 | 
  1596  | 
  note carr[simp] = factors_closed[OF afs ascarr] ascarr[THEN subsetD]
  | 
| 
 | 
  1597  | 
                    factors_closed[OF bfs bscarr] bscarr[THEN subsetD]
  | 
| 
 | 
  1598  | 
  | 
| 
 | 
  1599  | 
  from ab carr
  | 
| 
 | 
  1600  | 
      have "\<exists>u\<in>Units G. a = b \<otimes> u" by (fast elim: associatedE2)
  | 
| 
 | 
  1601  | 
  from this obtain u
  | 
| 
 | 
  1602  | 
      where uunit: "u \<in> Units G"
  | 
| 
 | 
  1603  | 
      and a: "a = b \<otimes> u" by auto
  | 
| 
 | 
  1604  | 
  | 
| 
 | 
  1605  | 
  from uunit bscarr
  | 
| 
 | 
  1606  | 
      have ee: "essentially_equal G (bs[0 := (bs!0 \<otimes> u)]) bs" 
  | 
| 
 | 
  1607  | 
                (is "essentially_equal G ?bs' bs")
  | 
| 
 | 
  1608  | 
      by (rule unitfactor_ee)
  | 
| 
 | 
  1609  | 
  | 
| 
 | 
  1610  | 
  from bscarr uunit
  | 
| 
 | 
  1611  | 
      have bs'carr: "set ?bs' \<subseteq> carrier G"
  | 
| 
 | 
  1612  | 
      by (cases bs) (simp add: Units_closed)+
  | 
| 
 | 
  1613  | 
  | 
| 
 | 
  1614  | 
  from uunit bnunit bfs bscarr
  | 
| 
 | 
  1615  | 
      have fac: "factors G ?bs' (b \<otimes> u)"
  | 
| 
 | 
  1616  | 
      by (rule factors_cong_unit)
  | 
| 
 | 
  1617  | 
  | 
| 
 | 
  1618  | 
  from afs fac[simplified a[symmetric]] ascarr bs'carr anunit
  | 
| 
 | 
  1619  | 
       have "essentially_equal G as ?bs'"
  | 
| 
 | 
  1620  | 
       by (blast intro: factors_unique)
  | 
| 
 | 
  1621  | 
  also note ee
  | 
| 
 | 
  1622  | 
  finally
  | 
| 
 | 
  1623  | 
      show "essentially_equal G as bs" by (simp add: ascarr bscarr bs'carr)
  | 
| 
 | 
  1624  | 
qed
  | 
| 
 | 
  1625  | 
  | 
| 
 | 
  1626  | 
lemma (in factorial_monoid) ee_wfactorsI:
  | 
| 
 | 
  1627  | 
  assumes asc: "a \<sim> b"
  | 
| 
 | 
  1628  | 
    and asf: "wfactors G as a" and bsf: "wfactors G bs b"
  | 
| 
 | 
  1629  | 
    and acarr[simp]: "a \<in> carrier G" and bcarr[simp]: "b \<in> carrier G"
  | 
| 
 | 
  1630  | 
    and ascarr[simp]: "set as \<subseteq> carrier G" and bscarr[simp]: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1631  | 
  shows "essentially_equal G as bs"
  | 
| 
 | 
  1632  | 
using assms
  | 
| 
 | 
  1633  | 
proof (cases "a \<in> Units G")
  | 
| 
 | 
  1634  | 
  assume aunit: "a \<in> Units G"
  | 
| 
 | 
  1635  | 
  also note asc
  | 
| 
 | 
  1636  | 
  finally have bunit: "b \<in> Units G" by simp
  | 
| 
 | 
  1637  | 
  | 
| 
 | 
  1638  | 
  from aunit asf ascarr
  | 
| 
 | 
  1639  | 
      have e: "as = []" by (rule unit_wfactors_empty)
  | 
| 
 | 
  1640  | 
  from bunit bsf bscarr
  | 
| 
 | 
  1641  | 
      have e': "bs = []" by (rule unit_wfactors_empty)
  | 
| 
 | 
  1642  | 
  | 
| 
 | 
  1643  | 
  have "essentially_equal G [] []"
  | 
| 
 | 
  1644  | 
      by (fast intro: essentially_equalI)
  | 
| 
 | 
  1645  | 
  thus ?thesis by (simp add: e e')
  | 
| 
 | 
  1646  | 
next
  | 
| 
 | 
  1647  | 
  assume anunit: "a \<notin> Units G"
  | 
| 
 | 
  1648  | 
  have bnunit: "b \<notin> Units G"
  | 
| 
 | 
  1649  | 
  proof clarify
  | 
| 
 | 
  1650  | 
    assume "b \<in> Units G"
  | 
| 
 | 
  1651  | 
    also note asc[symmetric]
  | 
| 
 | 
  1652  | 
    finally have "a \<in> Units G" by simp
  | 
| 
 | 
  1653  | 
    with anunit
  | 
| 
 | 
  1654  | 
         show "False" ..
  | 
| 
 | 
  1655  | 
  qed
  | 
| 
 | 
  1656  | 
  | 
| 
 | 
  1657  | 
  have "\<exists>a'. factors G as a' \<and> a' \<sim> a" by (rule wfactors_factors[OF asf ascarr])
  | 
| 
 | 
  1658  | 
  from this obtain a'
  | 
| 
 | 
  1659  | 
      where fa': "factors G as a'"
  | 
| 
 | 
  1660  | 
      and a': "a' \<sim> a"
  | 
| 
 | 
  1661  | 
      by auto
  | 
| 
 | 
  1662  | 
  from fa' ascarr
  | 
| 
 | 
  1663  | 
      have a'carr[simp]: "a' \<in> carrier G" by fast
  | 
| 
 | 
  1664  | 
  | 
| 
 | 
  1665  | 
  have a'nunit: "a' \<notin> Units G"
  | 
| 
 | 
  1666  | 
  proof (clarify)
  | 
| 
 | 
  1667  | 
    assume "a' \<in> Units G"
  | 
| 
 | 
  1668  | 
    also note a'
  | 
| 
 | 
  1669  | 
    finally have "a \<in> Units G" by simp
  | 
| 
 | 
  1670  | 
    with anunit
  | 
| 
 | 
  1671  | 
         show "False" ..
  | 
| 
 | 
  1672  | 
  qed
  | 
| 
 | 
  1673  | 
  | 
| 
 | 
  1674  | 
  have "\<exists>b'. factors G bs b' \<and> b' \<sim> b" by (rule wfactors_factors[OF bsf bscarr])
  | 
| 
 | 
  1675  | 
  from this obtain b'
  | 
| 
 | 
  1676  | 
      where fb': "factors G bs b'"
  | 
| 
 | 
  1677  | 
      and b': "b' \<sim> b"
  | 
| 
 | 
  1678  | 
      by auto
  | 
| 
 | 
  1679  | 
  from fb' bscarr
  | 
| 
 | 
  1680  | 
      have b'carr[simp]: "b' \<in> carrier G" by fast
  | 
| 
 | 
  1681  | 
  | 
| 
 | 
  1682  | 
  have b'nunit: "b' \<notin> Units G"
  | 
| 
 | 
  1683  | 
  proof (clarify)
  | 
| 
 | 
  1684  | 
    assume "b' \<in> Units G"
  | 
| 
 | 
  1685  | 
    also note b'
  | 
| 
 | 
  1686  | 
    finally have "b \<in> Units G" by simp
  | 
| 
 | 
  1687  | 
    with bnunit
  | 
| 
 | 
  1688  | 
        show "False" ..
  | 
| 
 | 
  1689  | 
  qed
  | 
| 
 | 
  1690  | 
  | 
| 
 | 
  1691  | 
  note a'
  | 
| 
 | 
  1692  | 
  also note asc
  | 
| 
 | 
  1693  | 
  also note b'[symmetric]
  | 
| 
 | 
  1694  | 
  finally
  | 
| 
 | 
  1695  | 
       have "a' \<sim> b'" by simp
  | 
| 
 | 
  1696  | 
  | 
| 
 | 
  1697  | 
  from this fa' a'nunit fb' b'nunit ascarr bscarr
  | 
| 
 | 
  1698  | 
  show "essentially_equal G as bs"
  | 
| 
 | 
  1699  | 
      by (rule ee_factorsI)
  | 
| 
 | 
  1700  | 
qed
  | 
| 
 | 
  1701  | 
  | 
| 
 | 
  1702  | 
lemma (in factorial_monoid) ee_wfactors:
  | 
| 
 | 
  1703  | 
  assumes asf: "wfactors G as a"
  | 
| 
 | 
  1704  | 
    and bsf: "wfactors G bs b"
  | 
| 
 | 
  1705  | 
    and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  1706  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1707  | 
  shows asc: "a \<sim> b = essentially_equal G as bs"
  | 
| 
 | 
  1708  | 
using assms
  | 
| 
 | 
  1709  | 
by (fast intro: ee_wfactorsI ee_wfactorsD)
  | 
| 
 | 
  1710  | 
  | 
| 
 | 
  1711  | 
lemma (in factorial_monoid) wfactors_exist [intro, simp]:
  | 
| 
 | 
  1712  | 
  assumes acarr[simp]: "a \<in> carrier G"
  | 
| 
 | 
  1713  | 
  shows "\<exists>fs. set fs \<subseteq> carrier G \<and> wfactors G fs a"
  | 
| 
 | 
  1714  | 
proof (cases "a \<in> Units G")
  | 
| 
 | 
  1715  | 
  assume "a \<in> Units G"
  | 
| 
 | 
  1716  | 
  hence "wfactors G [] a" by (rule unit_wfactors)
  | 
| 
 | 
  1717  | 
  thus ?thesis by (intro exI) force
  | 
| 
 | 
  1718  | 
next
  | 
| 
 | 
  1719  | 
  assume "a \<notin> Units G"
  | 
| 
 | 
  1720  | 
  hence "\<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs a" by (intro factors_exist acarr)
  | 
| 
 | 
  1721  | 
  from this obtain fs
  | 
| 
 | 
  1722  | 
      where fscarr: "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1723  | 
      and f: "factors G fs a"
  | 
| 
 | 
  1724  | 
      by auto
  | 
| 
 | 
  1725  | 
  from f have "wfactors G fs a" by (rule factors_wfactors) fact
  | 
| 
 | 
  1726  | 
  from fscarr this
  | 
| 
 | 
  1727  | 
      show ?thesis by fast
  | 
| 
 | 
  1728  | 
qed
  | 
| 
 | 
  1729  | 
  | 
| 
 | 
  1730  | 
lemma (in monoid) wfactors_prod_exists [intro, simp]:
  | 
| 
 | 
  1731  | 
  assumes "\<forall>a \<in> set as. irreducible G a" and "set as \<subseteq> carrier G"
  | 
| 
 | 
  1732  | 
  shows "\<exists>a. a \<in> carrier G \<and> wfactors G as a"
  | 
| 
 | 
  1733  | 
unfolding wfactors_def
  | 
| 
 | 
  1734  | 
using assms
  | 
| 
 | 
  1735  | 
by blast
  | 
| 
 | 
  1736  | 
  | 
| 
 | 
  1737  | 
lemma (in factorial_monoid) wfactors_unique:
  | 
| 
 | 
  1738  | 
  assumes "wfactors G fs a" and "wfactors G fs' a"
  | 
| 
 | 
  1739  | 
    and "a \<in> carrier G"
  | 
| 
 | 
  1740  | 
    and "set fs \<subseteq> carrier G" and "set fs' \<subseteq> carrier G"
  | 
| 
 | 
  1741  | 
  shows "essentially_equal G fs fs'"
  | 
| 
 | 
  1742  | 
using assms
  | 
| 
 | 
  1743  | 
by (fast intro: ee_wfactorsI[of a a])
  | 
| 
 | 
  1744  | 
  | 
| 
 | 
  1745  | 
lemma (in monoid) factors_mult_single:
  | 
| 
 | 
  1746  | 
  assumes "irreducible G a" and "factors G fb b" and "a \<in> carrier G"
  | 
| 
 | 
  1747  | 
  shows "factors G (a # fb) (a \<otimes> b)"
  | 
| 
 | 
  1748  | 
using assms
  | 
| 
 | 
  1749  | 
unfolding factors_def
  | 
| 
 | 
  1750  | 
by simp
  | 
| 
 | 
  1751  | 
  | 
| 
 | 
  1752  | 
lemma (in monoid_cancel) wfactors_mult_single:
  | 
| 
 | 
  1753  | 
  assumes f: "irreducible G a"  "wfactors G fb b"
  | 
| 
 | 
  1754  | 
        "a \<in> carrier G"  "b \<in> carrier G"  "set fb \<subseteq> carrier G"
  | 
| 
 | 
  1755  | 
  shows "wfactors G (a # fb) (a \<otimes> b)"
  | 
| 
 | 
  1756  | 
using assms
  | 
| 
 | 
  1757  | 
unfolding wfactors_def
  | 
| 
 | 
  1758  | 
by (simp add: mult_cong_r)
  | 
| 
 | 
  1759  | 
  | 
| 
 | 
  1760  | 
lemma (in monoid) factors_mult:
  | 
| 
 | 
  1761  | 
  assumes factors: "factors G fa a"  "factors G fb b"
  | 
| 
 | 
  1762  | 
    and ascarr: "set fa \<subseteq> carrier G" and bscarr:"set fb \<subseteq> carrier G"
  | 
| 
 | 
  1763  | 
  shows "factors G (fa @ fb) (a \<otimes> b)"
  | 
| 
 | 
  1764  | 
using assms
  | 
| 
 | 
  1765  | 
unfolding factors_def
  | 
| 
 | 
  1766  | 
apply (safe, force)
  | 
| 
 | 
  1767  | 
apply (induct fa)
  | 
| 
 | 
  1768  | 
 apply simp
  | 
| 
 | 
  1769  | 
apply (simp add: m_assoc)
  | 
| 
 | 
  1770  | 
done
  | 
| 
 | 
  1771  | 
  | 
| 
 | 
  1772  | 
lemma (in comm_monoid_cancel) wfactors_mult [intro]:
  | 
| 
 | 
  1773  | 
  assumes asf: "wfactors G as a" and bsf:"wfactors G bs b"
  | 
| 
 | 
  1774  | 
    and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  1775  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr:"set bs \<subseteq> carrier G"
  | 
| 
 | 
  1776  | 
  shows "wfactors G (as @ bs) (a \<otimes> b)"
  | 
| 
 | 
  1777  | 
apply (insert wfactors_factors[OF asf ascarr])
  | 
| 
 | 
  1778  | 
apply (insert wfactors_factors[OF bsf bscarr])
  | 
| 
 | 
  1779  | 
proof (clarsimp)
  | 
| 
 | 
  1780  | 
  fix a' b'
  | 
| 
 | 
  1781  | 
  assume asf': "factors G as a'" and a'a: "a' \<sim> a"
  | 
| 
 | 
  1782  | 
     and bsf': "factors G bs b'" and b'b: "b' \<sim> b"
  | 
| 
 | 
  1783  | 
  from asf' have a'carr: "a' \<in> carrier G" by (rule factors_closed) fact
  | 
| 
 | 
  1784  | 
  from bsf' have b'carr: "b' \<in> carrier G" by (rule factors_closed) fact
  | 
| 
 | 
  1785  | 
  | 
| 
 | 
  1786  | 
  note carr = acarr bcarr a'carr b'carr ascarr bscarr
  | 
| 
 | 
  1787  | 
  | 
| 
 | 
  1788  | 
  from asf' bsf'
  | 
| 
 | 
  1789  | 
      have "factors G (as @ bs) (a' \<otimes> b')" by (rule factors_mult) fact+
  | 
| 
 | 
  1790  | 
  | 
| 
 | 
  1791  | 
  with carr
  | 
| 
 | 
  1792  | 
       have abf': "wfactors G (as @ bs) (a' \<otimes> b')" by (intro factors_wfactors) simp+
  | 
| 
 | 
  1793  | 
  also from b'b carr
  | 
| 
 | 
  1794  | 
       have trb: "a' \<otimes> b' \<sim> a' \<otimes> b" by (intro mult_cong_r)
  | 
| 
 | 
  1795  | 
  also from a'a carr
  | 
| 
 | 
  1796  | 
       have tra: "a' \<otimes> b \<sim> a \<otimes> b" by (intro mult_cong_l)
  | 
| 
 | 
  1797  | 
  finally
  | 
| 
 | 
  1798  | 
       show "wfactors G (as @ bs) (a \<otimes> b)"
  | 
| 
 | 
  1799  | 
       by (simp add: carr)
  | 
| 
 | 
  1800  | 
qed
  | 
| 
 | 
  1801  | 
  | 
| 
 | 
  1802  | 
lemma (in comm_monoid) factors_dividesI:
  | 
| 
 | 
  1803  | 
  assumes "factors G fs a" and "f \<in> set fs"
  | 
| 
 | 
  1804  | 
    and "set fs \<subseteq> carrier G"
  | 
| 
 | 
  1805  | 
  shows "f divides a"
  | 
| 
 | 
  1806  | 
using assms
  | 
| 
 | 
  1807  | 
by (fast elim: factorsE intro: multlist_dividesI)
  | 
| 
 | 
  1808  | 
  | 
| 
 | 
  1809  | 
lemma (in comm_monoid) wfactors_dividesI:
  | 
| 
 | 
  1810  | 
  assumes p: "wfactors G fs a"
  | 
| 
 | 
  1811  | 
    and fscarr: "set fs \<subseteq> carrier G" and acarr: "a \<in> carrier G"
  | 
| 
 | 
  1812  | 
    and f: "f \<in> set fs"
  | 
| 
 | 
  1813  | 
  shows "f divides a"
  | 
| 
 | 
  1814  | 
apply (insert wfactors_factors[OF p fscarr], clarsimp)
  | 
| 
 | 
  1815  | 
proof -
  | 
| 
 | 
  1816  | 
  fix a'
  | 
| 
 | 
  1817  | 
  assume fsa': "factors G fs a'"
  | 
| 
 | 
  1818  | 
    and a'a: "a' \<sim> a"
  | 
| 
 | 
  1819  | 
  with fscarr
  | 
| 
 | 
  1820  | 
      have a'carr: "a' \<in> carrier G" by (simp add: factors_closed)
  | 
| 
 | 
  1821  | 
  | 
| 
 | 
  1822  | 
  from fsa' fscarr f
  | 
| 
 | 
  1823  | 
       have "f divides a'" by (fast intro: factors_dividesI)
  | 
| 
 | 
  1824  | 
  also note a'a
  | 
| 
 | 
  1825  | 
  finally
  | 
| 
 | 
  1826  | 
       show "f divides a" by (simp add: f fscarr[THEN subsetD] acarr a'carr)
  | 
| 
 | 
  1827  | 
qed
  | 
| 
 | 
  1828  | 
  | 
| 
 | 
  1829  | 
  | 
| 
 | 
  1830  | 
subsubsection {* Factorial monoids and wfactors *}
 | 
| 
 | 
  1831  | 
  | 
| 
 | 
  1832  | 
lemma (in comm_monoid_cancel) factorial_monoidI:
  | 
| 
 | 
  1833  | 
  assumes wfactors_exists: 
  | 
| 
 | 
  1834  | 
          "\<And>a. a \<in> carrier G \<Longrightarrow> \<exists>fs. set fs \<subseteq> carrier G \<and> wfactors G fs a"
  | 
| 
 | 
  1835  | 
      and wfactors_unique: 
  | 
| 
 | 
  1836  | 
          "\<And>a fs fs'. \<lbrakk>a \<in> carrier G; set fs \<subseteq> carrier G; set fs' \<subseteq> carrier G; 
  | 
| 
 | 
  1837  | 
                       wfactors G fs a; wfactors G fs' a\<rbrakk> \<Longrightarrow> essentially_equal G fs fs'"
  | 
| 
 | 
  1838  | 
  shows "factorial_monoid G"
  | 
| 
 | 
  1839  | 
proof (unfold_locales)
  | 
| 
 | 
  1840  | 
  fix a
  | 
| 
 | 
  1841  | 
  assume acarr: "a \<in> carrier G" and anunit: "a \<notin> Units G"
  | 
| 
 | 
  1842  | 
  | 
| 
 | 
  1843  | 
  from wfactors_exists[OF acarr]
  | 
| 
 | 
  1844  | 
  obtain as
  | 
| 
 | 
  1845  | 
      where ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  1846  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  1847  | 
      by auto
  | 
| 
 | 
  1848  | 
  from afs ascarr
  | 
| 
 | 
  1849  | 
      have "\<exists>a'. factors G as a' \<and> a' \<sim> a" by (rule wfactors_factors)
  | 
| 
 | 
  1850  | 
  from this obtain a'
  | 
| 
 | 
  1851  | 
      where afs': "factors G as a'"
  | 
| 
 | 
  1852  | 
      and a'a: "a' \<sim> a"
  | 
| 
 | 
  1853  | 
      by auto
  | 
| 
 | 
  1854  | 
  from afs' ascarr
  | 
| 
 | 
  1855  | 
      have a'carr: "a' \<in> carrier G" by fast
  | 
| 
 | 
  1856  | 
  have a'nunit: "a' \<notin> Units G"
  | 
| 
 | 
  1857  | 
  proof clarify
  | 
| 
 | 
  1858  | 
    assume "a' \<in> Units G"
  | 
| 
 | 
  1859  | 
    also note a'a
  | 
| 
 | 
  1860  | 
    finally have "a \<in> Units G" by (simp add: acarr)
  | 
| 
 | 
  1861  | 
    with anunit
  | 
| 
 | 
  1862  | 
        show "False" ..
  | 
| 
 | 
  1863  | 
  qed
  | 
| 
 | 
  1864  | 
  | 
| 
 | 
  1865  | 
  from a'carr acarr a'a
  | 
| 
 | 
  1866  | 
      have "\<exists>u. u \<in> Units G \<and> a' = a \<otimes> u" by (blast elim: associatedE2)
  | 
| 
 | 
  1867  | 
  from this obtain  u
  | 
| 
 | 
  1868  | 
      where uunit: "u \<in> Units G"
  | 
| 
 | 
  1869  | 
      and a': "a' = a \<otimes> u"
  | 
| 
 | 
  1870  | 
      by auto
  | 
| 
 | 
  1871  | 
  | 
| 
 | 
  1872  | 
  note [simp] = acarr Units_closed[OF uunit] Units_inv_closed[OF uunit]
  | 
| 
 | 
  1873  | 
  | 
| 
 | 
  1874  | 
  have "a = a \<otimes> \<one>" by simp
  | 
| 
 | 
  1875  | 
  also have "\<dots> = a \<otimes> (u \<otimes> inv u)" by (simp add: Units_r_inv uunit)
  | 
| 
 | 
  1876  | 
  also have "\<dots> = a' \<otimes> inv u" by (simp add: m_assoc[symmetric] a'[symmetric])
  | 
| 
 | 
  1877  | 
  finally
  | 
| 
 | 
  1878  | 
       have a: "a = a' \<otimes> inv u" .
  | 
| 
 | 
  1879  | 
  | 
| 
 | 
  1880  | 
  from ascarr uunit
  | 
| 
 | 
  1881  | 
      have cr: "set (as[0:=(as!0 \<otimes> inv u)]) \<subseteq> carrier G"
  | 
| 
 | 
  1882  | 
      by (cases as, clarsimp+)
  | 
| 
 | 
  1883  | 
  | 
| 
 | 
  1884  | 
  from afs' uunit a'nunit acarr ascarr
  | 
| 
 | 
  1885  | 
      have "factors G (as[0:=(as!0 \<otimes> inv u)]) a"
  | 
| 
 | 
  1886  | 
      by (simp add: a factors_cong_unit)
  | 
| 
 | 
  1887  | 
  | 
| 
 | 
  1888  | 
  with cr
  | 
| 
 | 
  1889  | 
      show "\<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs a" by fast
  | 
| 
 | 
  1890  | 
qed (blast intro: factors_wfactors wfactors_unique)
  | 
| 
 | 
  1891  | 
  | 
| 
 | 
  1892  | 
  | 
| 
 | 
  1893  | 
subsection {* Factorizations as multisets *}
 | 
| 
 | 
  1894  | 
  | 
| 
 | 
  1895  | 
text {* Gives useful operations like intersection *}
 | 
| 
 | 
  1896  | 
  | 
| 
 | 
  1897  | 
(* FIXME: use class_of x instead of closure_of {x} *)
 | 
| 
 | 
  1898  | 
  | 
| 
 | 
  1899  | 
abbreviation
  | 
| 
 | 
  1900  | 
  "assocs G x == eq_closure_of (division_rel G) {x}"
 | 
| 
 | 
  1901  | 
  | 
| 
 | 
  1902  | 
constdefs (structure G)
  | 
| 
 | 
  1903  | 
  "fmset G as \<equiv> multiset_of (map (\<lambda>a. assocs G a) as)"
  | 
| 
 | 
  1904  | 
  | 
| 
 | 
  1905  | 
  | 
| 
 | 
  1906  | 
text {* Helper lemmas *}
 | 
| 
 | 
  1907  | 
  | 
| 
 | 
  1908  | 
lemma (in monoid) assocs_repr_independence:
  | 
| 
 | 
  1909  | 
  assumes "y \<in> assocs G x"
  | 
| 
 | 
  1910  | 
    and "x \<in> carrier G"
  | 
| 
 | 
  1911  | 
  shows "assocs G x = assocs G y"
  | 
| 
 | 
  1912  | 
using assms
  | 
| 
 | 
  1913  | 
apply safe
  | 
| 
 | 
  1914  | 
 apply (elim closure_ofE2, intro closure_ofI2[of _ _ y])
  | 
| 
 | 
  1915  | 
   apply (clarsimp, iprover intro: associated_trans associated_sym, simp+)
  | 
| 
 | 
  1916  | 
apply (elim closure_ofE2, intro closure_ofI2[of _ _ x])
  | 
| 
 | 
  1917  | 
  apply (clarsimp, iprover intro: associated_trans, simp+)
  | 
| 
 | 
  1918  | 
done
  | 
| 
 | 
  1919  | 
  | 
| 
 | 
  1920  | 
lemma (in monoid) assocs_self:
  | 
| 
 | 
  1921  | 
  assumes "x \<in> carrier G"
  | 
| 
 | 
  1922  | 
  shows "x \<in> assocs G x"
  | 
| 
 | 
  1923  | 
using assms
  | 
| 
 | 
  1924  | 
by (fastsimp intro: closure_ofI2)
  | 
| 
 | 
  1925  | 
  | 
| 
 | 
  1926  | 
lemma (in monoid) assocs_repr_independenceD:
  | 
| 
 | 
  1927  | 
  assumes repr: "assocs G x = assocs G y"
  | 
| 
 | 
  1928  | 
    and ycarr: "y \<in> carrier G"
  | 
| 
 | 
  1929  | 
  shows "y \<in> assocs G x"
  | 
| 
 | 
  1930  | 
unfolding repr
  | 
| 
 | 
  1931  | 
using ycarr
  | 
| 
 | 
  1932  | 
by (intro assocs_self)
  | 
| 
 | 
  1933  | 
  | 
| 
 | 
  1934  | 
lemma (in comm_monoid) assocs_assoc:
  | 
| 
 | 
  1935  | 
  assumes "a \<in> assocs G b"
  | 
| 
 | 
  1936  | 
    and "b \<in> carrier G"
  | 
| 
 | 
  1937  | 
  shows "a \<sim> b"
  | 
| 
 | 
  1938  | 
using assms
  | 
| 
 | 
  1939  | 
by (elim closure_ofE2, simp)
  | 
| 
 | 
  1940  | 
  | 
| 
 | 
  1941  | 
lemmas (in comm_monoid) assocs_eqD =
  | 
| 
 | 
  1942  | 
    assocs_repr_independenceD[THEN assocs_assoc]
  | 
| 
 | 
  1943  | 
  | 
| 
 | 
  1944  | 
  | 
| 
 | 
  1945  | 
subsubsection {* Comparing multisets *}
 | 
| 
 | 
  1946  | 
  | 
| 
 | 
  1947  | 
lemma (in monoid) fmset_perm_cong:
  | 
| 
 | 
  1948  | 
  assumes prm: "as <~~> bs"
  | 
| 
 | 
  1949  | 
  shows "fmset G as = fmset G bs"
  | 
| 
 | 
  1950  | 
using perm_map[OF prm]
  | 
| 
 | 
  1951  | 
by (simp add: multiset_of_eq_perm fmset_def)
  | 
| 
 | 
  1952  | 
  | 
| 
 | 
  1953  | 
lemma (in comm_monoid_cancel) eqc_listassoc_cong:
  | 
| 
 | 
  1954  | 
  assumes "as [\<sim>] bs"
  | 
| 
 | 
  1955  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1956  | 
  shows "map (assocs G) as = map (assocs G) bs"
  | 
| 
 | 
  1957  | 
using assms
  | 
| 
 | 
  1958  | 
apply (induct as arbitrary: bs, simp)
  | 
| 
 | 
  1959  | 
apply (clarsimp simp add: Cons_eq_map_conv list_all2_Cons1, safe)
  | 
| 
 | 
  1960  | 
 apply (clarsimp elim!: closure_ofE2) defer 1
  | 
| 
 | 
  1961  | 
 apply (clarsimp elim!: closure_ofE2) defer 1
  | 
| 
 | 
  1962  | 
proof -
  | 
| 
 | 
  1963  | 
  fix a x z
  | 
| 
 | 
  1964  | 
  assume carr[simp]: "a \<in> carrier G"  "x \<in> carrier G"  "z \<in> carrier G"
  | 
| 
 | 
  1965  | 
  assume "x \<sim> a"
  | 
| 
 | 
  1966  | 
  also assume "a \<sim> z"
  | 
| 
 | 
  1967  | 
  finally have "x \<sim> z" by simp
  | 
| 
 | 
  1968  | 
  with carr
  | 
| 
 | 
  1969  | 
      show "x \<in> assocs G z"
  | 
| 
 | 
  1970  | 
      by (intro closure_ofI2) simp+
  | 
| 
 | 
  1971  | 
next
  | 
| 
 | 
  1972  | 
  fix a x z
  | 
| 
 | 
  1973  | 
  assume carr[simp]: "a \<in> carrier G"  "x \<in> carrier G"  "z \<in> carrier G"
  | 
| 
 | 
  1974  | 
  assume "x \<sim> z"
  | 
| 
 | 
  1975  | 
  also assume [symmetric]: "a \<sim> z"
  | 
| 
 | 
  1976  | 
  finally have "x \<sim> a" by simp
  | 
| 
 | 
  1977  | 
  with carr
  | 
| 
 | 
  1978  | 
      show "x \<in> assocs G a"
  | 
| 
 | 
  1979  | 
      by (intro closure_ofI2) simp+
  | 
| 
 | 
  1980  | 
qed
  | 
| 
 | 
  1981  | 
  | 
| 
 | 
  1982  | 
lemma (in comm_monoid_cancel) fmset_listassoc_cong:
  | 
| 
 | 
  1983  | 
  assumes "as [\<sim>] bs" 
  | 
| 
 | 
  1984  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1985  | 
  shows "fmset G as = fmset G bs"
  | 
| 
 | 
  1986  | 
using assms
  | 
| 
 | 
  1987  | 
unfolding fmset_def
  | 
| 
 | 
  1988  | 
by (simp add: eqc_listassoc_cong)
  | 
| 
 | 
  1989  | 
  | 
| 
 | 
  1990  | 
lemma (in comm_monoid_cancel) ee_fmset:
  | 
| 
 | 
  1991  | 
  assumes ee: "essentially_equal G as bs" 
  | 
| 
 | 
  1992  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  1993  | 
  shows "fmset G as = fmset G bs"
  | 
| 
 | 
  1994  | 
using ee
  | 
| 
 | 
  1995  | 
proof (elim essentially_equalE)
  | 
| 
 | 
  1996  | 
  fix as'
  | 
| 
 | 
  1997  | 
  assume prm: "as <~~> as'"
  | 
| 
 | 
  1998  | 
  from prm ascarr
  | 
| 
 | 
  1999  | 
      have as'carr: "set as' \<subseteq> carrier G" by (rule perm_closed)
  | 
| 
 | 
  2000  | 
  | 
| 
 | 
  2001  | 
  from prm
  | 
| 
 | 
  2002  | 
       have "fmset G as = fmset G as'" by (rule fmset_perm_cong)
  | 
| 
 | 
  2003  | 
  also assume "as' [\<sim>] bs"
  | 
| 
 | 
  2004  | 
       with as'carr bscarr
  | 
| 
 | 
  2005  | 
       have "fmset G as' = fmset G bs" by (simp add: fmset_listassoc_cong)
  | 
| 
 | 
  2006  | 
  finally
  | 
| 
 | 
  2007  | 
       show "fmset G as = fmset G bs" .
  | 
| 
 | 
  2008  | 
qed
  | 
| 
 | 
  2009  | 
  | 
| 
 | 
  2010  | 
lemma (in monoid_cancel) fmset_ee__hlp_induct:
  | 
| 
 | 
  2011  | 
  assumes prm: "cas <~~> cbs"
  | 
| 
 | 
  2012  | 
    and cdef: "cas = map (assocs G) as"  "cbs = map (assocs G) bs"
  | 
| 
 | 
  2013  | 
  shows "\<forall>as bs. (cas <~~> cbs \<and> cas = map (assocs G) as \<and> 
  | 
| 
 | 
  2014  | 
                 cbs = map (assocs G) bs) \<longrightarrow> (\<exists>as'. as <~~> as' \<and> map (assocs G) as' = cbs)"
  | 
| 
 | 
  2015  | 
apply (rule perm.induct[of cas cbs], rule prm)
  | 
| 
 | 
  2016  | 
apply safe apply simp_all
  | 
| 
 | 
  2017  | 
  apply (simp add: map_eq_Cons_conv, blast)
  | 
| 
 | 
  2018  | 
 apply force
  | 
| 
 | 
  2019  | 
proof -
  | 
| 
 | 
  2020  | 
  fix ys as bs
  | 
| 
 | 
  2021  | 
  assume p1: "map (assocs G) as <~~> ys"
  | 
| 
 | 
  2022  | 
    and r1[rule_format]:
  | 
| 
 | 
  2023  | 
        "\<forall>asa bs. map (assocs G) as = map (assocs G) asa \<and>
  | 
| 
 | 
  2024  | 
                  ys = map (assocs G) bs
  | 
| 
 | 
  2025  | 
                  \<longrightarrow> (\<exists>as'. asa <~~> as' \<and> map (assocs G) as' = map (assocs G) bs)"
  | 
| 
 | 
  2026  | 
    and p2: "ys <~~> map (assocs G) bs"
  | 
| 
 | 
  2027  | 
    and r2[rule_format]:
  | 
| 
 | 
  2028  | 
        "\<forall>as bsa. ys = map (assocs G) as \<and>
  | 
| 
 | 
  2029  | 
                  map (assocs G) bs = map (assocs G) bsa
  | 
| 
 | 
  2030  | 
                  \<longrightarrow> (\<exists>as'. as <~~> as' \<and> map (assocs G) as' = map (assocs G) bsa)"
  | 
| 
 | 
  2031  | 
    and p3: "map (assocs G) as <~~> map (assocs G) bs"
  | 
| 
 | 
  2032  | 
  | 
| 
 | 
  2033  | 
  from p1
  | 
| 
 | 
  2034  | 
      have "multiset_of (map (assocs G) as) = multiset_of ys"
  | 
| 
 | 
  2035  | 
      by (simp add: multiset_of_eq_perm)
  | 
| 
 | 
  2036  | 
  hence setys: "set (map (assocs G) as) = set ys" by (rule multiset_of_eq_setD)
  | 
| 
 | 
  2037  | 
  | 
| 
 | 
  2038  | 
  have "set (map (assocs G) as) = { assocs G x | x. x \<in> set as}" by clarsimp fast
 | 
| 
 | 
  2039  | 
  with setys have "set ys \<subseteq> { assocs G x | x. x \<in> set as}" by simp
 | 
| 
 | 
  2040  | 
  hence "\<exists>yy. ys = map (assocs G) yy"
  | 
| 
 | 
  2041  | 
    apply (induct ys, simp, clarsimp)
  | 
| 
 | 
  2042  | 
  proof -
  | 
| 
 | 
  2043  | 
    fix yy x
  | 
| 
 | 
  2044  | 
    show "\<exists>yya. (assocs G x) # map (assocs G) yy =
  | 
| 
 | 
  2045  | 
                map (assocs G) yya"
  | 
| 
 | 
  2046  | 
    by (rule exI[of _ "x#yy"], simp)
  | 
| 
 | 
  2047  | 
  qed
  | 
| 
 | 
  2048  | 
  from this obtain yy
  | 
| 
 | 
  2049  | 
      where ys: "ys = map (assocs G) yy"
  | 
| 
 | 
  2050  | 
      by auto
  | 
| 
 | 
  2051  | 
  | 
| 
 | 
  2052  | 
  from p1 ys
  | 
| 
 | 
  2053  | 
      have "\<exists>as'. as <~~> as' \<and> map (assocs G) as' = map (assocs G) yy"
  | 
| 
 | 
  2054  | 
      by (intro r1, simp)
  | 
| 
 | 
  2055  | 
  from this obtain as'
  | 
| 
 | 
  2056  | 
      where asas': "as <~~> as'"
  | 
| 
 | 
  2057  | 
      and as'yy: "map (assocs G) as' = map (assocs G) yy"
  | 
| 
 | 
  2058  | 
      by auto
  | 
| 
 | 
  2059  | 
  | 
| 
 | 
  2060  | 
  from p2 ys
  | 
| 
 | 
  2061  | 
      have "\<exists>as'. yy <~~> as' \<and> map (assocs G) as' = map (assocs G) bs"
  | 
| 
 | 
  2062  | 
      by (intro r2, simp)
  | 
| 
 | 
  2063  | 
  from this obtain as''
  | 
| 
 | 
  2064  | 
      where yyas'': "yy <~~> as''"
  | 
| 
 | 
  2065  | 
      and as''bs: "map (assocs G) as'' = map (assocs G) bs"
  | 
| 
 | 
  2066  | 
      by auto
  | 
| 
 | 
  2067  | 
  | 
| 
 | 
  2068  | 
  from as'yy and yyas''
  | 
| 
 | 
  2069  | 
      have "\<exists>cs. as' <~~> cs \<and> map (assocs G) cs = map (assocs G) as''"
  | 
| 
 | 
  2070  | 
      by (rule perm_map_switch)
  | 
| 
 | 
  2071  | 
  from this obtain cs
  | 
| 
 | 
  2072  | 
      where as'cs: "as' <~~> cs"
  | 
| 
 | 
  2073  | 
      and csas'': "map (assocs G) cs = map (assocs G) as''"
  | 
| 
 | 
  2074  | 
      by auto
  | 
| 
 | 
  2075  | 
  | 
| 
 | 
  2076  | 
  from asas' and as'cs
  | 
| 
 | 
  2077  | 
      have ascs: "as <~~> cs" by fast
  | 
| 
 | 
  2078  | 
  from csas'' and as''bs
  | 
| 
 | 
  2079  | 
      have "map (assocs G) cs = map (assocs G) bs" by simp
  | 
| 
 | 
  2080  | 
  from ascs and this
  | 
| 
 | 
  2081  | 
  show "\<exists>as'. as <~~> as' \<and> map (assocs G) as' = map (assocs G) bs" by fast
  | 
| 
 | 
  2082  | 
qed
  | 
| 
 | 
  2083  | 
  | 
| 
 | 
  2084  | 
lemma (in comm_monoid_cancel) fmset_ee:
  | 
| 
 | 
  2085  | 
  assumes mset: "fmset G as = fmset G bs"
  | 
| 
 | 
  2086  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2087  | 
  shows "essentially_equal G as bs"
  | 
| 
 | 
  2088  | 
proof -
  | 
| 
 | 
  2089  | 
  from mset
  | 
| 
 | 
  2090  | 
      have mpp: "map (assocs G) as <~~> map (assocs G) bs"
  | 
| 
 | 
  2091  | 
      by (simp add: fmset_def multiset_of_eq_perm)
  | 
| 
 | 
  2092  | 
  | 
| 
 | 
  2093  | 
  have "\<exists>cas. cas = map (assocs G) as" by simp
  | 
| 
 | 
  2094  | 
  from this obtain cas where cas: "cas = map (assocs G) as" by simp
  | 
| 
 | 
  2095  | 
  | 
| 
 | 
  2096  | 
  have "\<exists>cbs. cbs = map (assocs G) bs" by simp
  | 
| 
 | 
  2097  | 
  from this obtain cbs where cbs: "cbs = map (assocs G) bs" by simp
  | 
| 
 | 
  2098  | 
  | 
| 
 | 
  2099  | 
  from cas cbs mpp
  | 
| 
 | 
  2100  | 
      have [rule_format]:
  | 
| 
 | 
  2101  | 
           "\<forall>as bs. (cas <~~> cbs \<and> cas = map (assocs G) as \<and> 
  | 
| 
 | 
  2102  | 
                     cbs = map (assocs G) bs) 
  | 
| 
 | 
  2103  | 
                     \<longrightarrow> (\<exists>as'. as <~~> as' \<and> map (assocs G) as' = cbs)"
  | 
| 
 | 
  2104  | 
      by (intro fmset_ee__hlp_induct, simp+)
  | 
| 
 | 
  2105  | 
  with mpp cas cbs
  | 
| 
 | 
  2106  | 
      have "\<exists>as'. as <~~> as' \<and> map (assocs G) as' = map (assocs G) bs"
  | 
| 
 | 
  2107  | 
      by simp
  | 
| 
 | 
  2108  | 
  | 
| 
 | 
  2109  | 
  from this obtain as'
  | 
| 
 | 
  2110  | 
      where tp: "as <~~> as'"
  | 
| 
 | 
  2111  | 
      and tm: "map (assocs G) as' = map (assocs G) bs"
  | 
| 
 | 
  2112  | 
      by auto
  | 
| 
 | 
  2113  | 
  from tm have lene: "length as' = length bs" by (rule map_eq_imp_length_eq)
  | 
| 
 | 
  2114  | 
  from tp have "set as = set as'" by (simp add: multiset_of_eq_perm multiset_of_eq_setD)
  | 
| 
 | 
  2115  | 
  with ascarr
  | 
| 
 | 
  2116  | 
      have as'carr: "set as' \<subseteq> carrier G" by simp
  | 
| 
 | 
  2117  | 
  | 
| 
 | 
  2118  | 
  from tm as'carr[THEN subsetD] bscarr[THEN subsetD]
  | 
| 
 | 
  2119  | 
  have "as' [\<sim>] bs"
  | 
| 
 | 
  2120  | 
    by (induct as' arbitrary: bs) (simp, fastsimp dest: assocs_eqD[THEN associated_sym])
  | 
| 
 | 
  2121  | 
  | 
| 
 | 
  2122  | 
  from tp and this
  | 
| 
 | 
  2123  | 
    show "essentially_equal G as bs" by (fast intro: essentially_equalI)
  | 
| 
 | 
  2124  | 
qed
  | 
| 
 | 
  2125  | 
  | 
| 
 | 
  2126  | 
lemma (in comm_monoid_cancel) ee_is_fmset:
  | 
| 
 | 
  2127  | 
  assumes "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2128  | 
  shows "essentially_equal G as bs = (fmset G as = fmset G bs)"
  | 
| 
 | 
  2129  | 
using assms
  | 
| 
 | 
  2130  | 
by (fast intro: ee_fmset fmset_ee)
  | 
| 
 | 
  2131  | 
  | 
| 
 | 
  2132  | 
  | 
| 
 | 
  2133  | 
subsubsection {* Interpreting multisets as factorizations *}
 | 
| 
 | 
  2134  | 
  | 
| 
 | 
  2135  | 
lemma (in monoid) mset_fmsetEx:
  | 
| 
 | 
  2136  | 
  assumes elems: "\<And>X. X \<in> set_of Cs \<Longrightarrow> \<exists>x. P x \<and> X = assocs G x"
  | 
| 
 | 
  2137  | 
  shows "\<exists>cs. (\<forall>c \<in> set cs. P c) \<and> fmset G cs = Cs"
  | 
| 
 | 
  2138  | 
proof -
  | 
| 
 | 
  2139  | 
  have "\<exists>Cs'. Cs = multiset_of Cs'"
  | 
| 
 | 
  2140  | 
      by (rule surjE[OF surj_multiset_of], fast)
  | 
| 
 | 
  2141  | 
  from this obtain Cs'
  | 
| 
 | 
  2142  | 
      where Cs: "Cs = multiset_of Cs'"
  | 
| 
 | 
  2143  | 
      by auto
  | 
| 
 | 
  2144  | 
  | 
| 
 | 
  2145  | 
  have "\<exists>cs. (\<forall>c \<in> set cs. P c) \<and> multiset_of (map (assocs G) cs) = Cs"
  | 
| 
 | 
  2146  | 
  using elems
  | 
| 
 | 
  2147  | 
  unfolding Cs
  | 
| 
 | 
  2148  | 
    apply (induct Cs', simp)
  | 
| 
 | 
  2149  | 
    apply clarsimp
  | 
| 
 | 
  2150  | 
    apply (subgoal_tac "\<exists>cs. (\<forall>x\<in>set cs. P x) \<and> 
  | 
| 
 | 
  2151  | 
                             multiset_of (map (assocs G) cs) = multiset_of Cs'")
  | 
| 
 | 
  2152  | 
  proof clarsimp
  | 
| 
 | 
  2153  | 
    fix a Cs' cs 
  | 
| 
 | 
  2154  | 
    assume ih: "\<And>X. X = a \<or> X \<in> set Cs' \<Longrightarrow> \<exists>x. P x \<and> X = assocs G x"
  | 
| 
 | 
  2155  | 
      and csP: "\<forall>x\<in>set cs. P x"
  | 
| 
 | 
  2156  | 
      and mset: "multiset_of (map (assocs G) cs) = multiset_of Cs'"
  | 
| 
 | 
  2157  | 
    from ih
  | 
| 
 | 
  2158  | 
        have "\<exists>x. P x \<and> a = assocs G x" by fast
  | 
| 
 | 
  2159  | 
    from this obtain c
  | 
| 
 | 
  2160  | 
        where cP: "P c"
  | 
| 
 | 
  2161  | 
        and a: "a = assocs G c"
  | 
| 
 | 
  2162  | 
        by auto
  | 
| 
 | 
  2163  | 
    from cP csP
  | 
| 
 | 
  2164  | 
        have tP: "\<forall>x\<in>set (c#cs). P x" by simp
  | 
| 
 | 
  2165  | 
    from mset a
  | 
| 
 | 
  2166  | 
    have "multiset_of (map (assocs G) (c#cs)) = multiset_of Cs' + {#a#}" by simp
 | 
| 
 | 
  2167  | 
    from tP this
  | 
| 
 | 
  2168  | 
    show "\<exists>cs. (\<forall>x\<in>set cs. P x) \<and>
  | 
| 
 | 
  2169  | 
               multiset_of (map (assocs G) cs) =
  | 
| 
 | 
  2170  | 
               multiset_of Cs' + {#a#}" by fast
 | 
| 
 | 
  2171  | 
  qed simp
  | 
| 
 | 
  2172  | 
  thus ?thesis by (simp add: fmset_def)
  | 
| 
 | 
  2173  | 
qed
  | 
| 
 | 
  2174  | 
  | 
| 
 | 
  2175  | 
lemma (in monoid) mset_wfactorsEx:
  | 
| 
 | 
  2176  | 
  assumes elems: "\<And>X. X \<in> set_of Cs 
  | 
| 
 | 
  2177  | 
                      \<Longrightarrow> \<exists>x. (x \<in> carrier G \<and> irreducible G x) \<and> X = assocs G x"
  | 
| 
 | 
  2178  | 
  shows "\<exists>c cs. c \<in> carrier G \<and> set cs \<subseteq> carrier G \<and> wfactors G cs c \<and> fmset G cs = Cs"
  | 
| 
 | 
  2179  | 
proof -
  | 
| 
 | 
  2180  | 
  have "\<exists>cs. (\<forall>c\<in>set cs. c \<in> carrier G \<and> irreducible G c) \<and> fmset G cs = Cs"
  | 
| 
 | 
  2181  | 
      by (intro mset_fmsetEx, rule elems)
  | 
| 
 | 
  2182  | 
  from this obtain cs
  | 
| 
 | 
  2183  | 
      where p[rule_format]: "\<forall>c\<in>set cs. c \<in> carrier G \<and> irreducible G c"
  | 
| 
 | 
  2184  | 
      and Cs[symmetric]: "fmset G cs = Cs"
  | 
| 
 | 
  2185  | 
      by auto
  | 
| 
 | 
  2186  | 
  | 
| 
 | 
  2187  | 
  from p
  | 
| 
 | 
  2188  | 
      have cscarr: "set cs \<subseteq> carrier G" by fast
  | 
| 
 | 
  2189  | 
  | 
| 
 | 
  2190  | 
  from p
  | 
| 
 | 
  2191  | 
      have "\<exists>c. c \<in> carrier G \<and> wfactors G cs c"
  | 
| 
 | 
  2192  | 
      by (intro wfactors_prod_exists) fast+
  | 
| 
 | 
  2193  | 
  from this obtain c
  | 
| 
 | 
  2194  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2195  | 
      and cfs: "wfactors G cs c"
  | 
| 
 | 
  2196  | 
      by auto
  | 
| 
 | 
  2197  | 
  | 
| 
 | 
  2198  | 
  with cscarr Cs
  | 
| 
 | 
  2199  | 
      show ?thesis by fast
  | 
| 
 | 
  2200  | 
qed
  | 
| 
 | 
  2201  | 
  | 
| 
 | 
  2202  | 
  | 
| 
 | 
  2203  | 
subsubsection {* Multiplication on multisets *}
 | 
| 
 | 
  2204  | 
  | 
| 
 | 
  2205  | 
lemma (in factorial_monoid) mult_wfactors_fmset:
  | 
| 
 | 
  2206  | 
  assumes afs: "wfactors G as a" and bfs: "wfactors G bs b" and cfs: "wfactors G cs (a \<otimes> b)"
  | 
| 
 | 
  2207  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2208  | 
              "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"  "set cs \<subseteq> carrier G"
  | 
| 
 | 
  2209  | 
  shows "fmset G cs = fmset G as + fmset G bs"
  | 
| 
 | 
  2210  | 
proof -
  | 
| 
 | 
  2211  | 
  from assms
  | 
| 
 | 
  2212  | 
       have "wfactors G (as @ bs) (a \<otimes> b)" by (intro wfactors_mult)
  | 
| 
 | 
  2213  | 
  with carr cfs
  | 
| 
 | 
  2214  | 
       have "essentially_equal G cs (as@bs)" by (intro ee_wfactorsI[of "a\<otimes>b" "a\<otimes>b"], simp+)
  | 
| 
 | 
  2215  | 
  with carr
  | 
| 
 | 
  2216  | 
       have "fmset G cs = fmset G (as@bs)" by (intro ee_fmset, simp+)
  | 
| 
 | 
  2217  | 
  also have "fmset G (as@bs) = fmset G as + fmset G bs" by (simp add: fmset_def)
  | 
| 
 | 
  2218  | 
  finally show "fmset G cs = fmset G as + fmset G bs" .
  | 
| 
 | 
  2219  | 
qed
  | 
| 
 | 
  2220  | 
  | 
| 
 | 
  2221  | 
lemma (in factorial_monoid) mult_factors_fmset:
  | 
| 
 | 
  2222  | 
  assumes afs: "factors G as a" and bfs: "factors G bs b" and cfs: "factors G cs (a \<otimes> b)"
  | 
| 
 | 
  2223  | 
    and "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"  "set cs \<subseteq> carrier G"
  | 
| 
 | 
  2224  | 
  shows "fmset G cs = fmset G as + fmset G bs"
  | 
| 
 | 
  2225  | 
using assms
  | 
| 
 | 
  2226  | 
by (blast intro: factors_wfactors mult_wfactors_fmset)
  | 
| 
 | 
  2227  | 
  | 
| 
 | 
  2228  | 
lemma (in comm_monoid_cancel) fmset_wfactors_mult:
  | 
| 
 | 
  2229  | 
  assumes mset: "fmset G cs = fmset G as + fmset G bs"
  | 
| 
 | 
  2230  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
  2231  | 
          "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"  "set cs \<subseteq> carrier G"
  | 
| 
 | 
  2232  | 
    and fs: "wfactors G as a"  "wfactors G bs b"  "wfactors G cs c"
  | 
| 
 | 
  2233  | 
  shows "c \<sim> a \<otimes> b"
  | 
| 
 | 
  2234  | 
proof -
  | 
| 
 | 
  2235  | 
  from carr fs
  | 
| 
 | 
  2236  | 
       have m: "wfactors G (as @ bs) (a \<otimes> b)" by (intro wfactors_mult)
  | 
| 
 | 
  2237  | 
  | 
| 
 | 
  2238  | 
  from mset
  | 
| 
 | 
  2239  | 
       have "fmset G cs = fmset G (as@bs)" by (simp add: fmset_def)
  | 
| 
 | 
  2240  | 
  then have "essentially_equal G cs (as@bs)" by (rule fmset_ee) (simp add: carr)+
  | 
| 
 | 
  2241  | 
  then show "c \<sim> a \<otimes> b" by (rule ee_wfactorsD[of "cs" "as@bs"]) (simp add: assms m)+
  | 
| 
 | 
  2242  | 
qed
  | 
| 
 | 
  2243  | 
  | 
| 
 | 
  2244  | 
  | 
| 
 | 
  2245  | 
subsubsection {* Divisibility on multisets *}
 | 
| 
 | 
  2246  | 
  | 
| 
 | 
  2247  | 
lemma (in factorial_monoid) divides_fmsubset:
  | 
| 
 | 
  2248  | 
  assumes ab: "a divides b"
  | 
| 
 | 
  2249  | 
    and afs: "wfactors G as a" and bfs: "wfactors G bs b"
  | 
| 
 | 
  2250  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"  "set as \<subseteq> carrier G"  "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2251  | 
  shows "fmset G as \<le># fmset G bs"
  | 
| 
 | 
  2252  | 
using ab
  | 
| 
 | 
  2253  | 
proof (elim dividesE)
  | 
| 
 | 
  2254  | 
  fix c
  | 
| 
 | 
  2255  | 
  assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2256  | 
  hence "\<exists>cs. set cs \<subseteq> carrier G \<and> wfactors G cs c" by (rule wfactors_exist)
  | 
| 
 | 
  2257  | 
  from this obtain cs 
  | 
| 
 | 
  2258  | 
      where cscarr: "set cs \<subseteq> carrier G"
  | 
| 
 | 
  2259  | 
      and cfs: "wfactors G cs c" by auto
  | 
| 
 | 
  2260  | 
  note carr = carr ccarr cscarr
  | 
| 
 | 
  2261  | 
  | 
| 
 | 
  2262  | 
  assume "b = a \<otimes> c"
  | 
| 
 | 
  2263  | 
  with afs bfs cfs carr
  | 
| 
 | 
  2264  | 
      have "fmset G bs = fmset G as + fmset G cs"
  | 
| 
 | 
  2265  | 
      by (intro mult_wfactors_fmset[OF afs cfs]) simp+
  | 
| 
 | 
  2266  | 
  | 
| 
 | 
  2267  | 
  thus ?thesis by simp
  | 
| 
 | 
  2268  | 
qed
  | 
| 
 | 
  2269  | 
  | 
| 
 | 
  2270  | 
lemma (in comm_monoid_cancel) fmsubset_divides:
  | 
| 
 | 
  2271  | 
  assumes msubset: "fmset G as \<le># fmset G bs"
  | 
| 
 | 
  2272  | 
    and afs: "wfactors G as a" and bfs: "wfactors G bs b"
  | 
| 
 | 
  2273  | 
    and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  2274  | 
    and ascarr: "set as \<subseteq> carrier G" and bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2275  | 
  shows "a divides b"
  | 
| 
 | 
  2276  | 
proof -
  | 
| 
 | 
  2277  | 
  from afs have airr: "\<forall>a \<in> set as. irreducible G a" by (fast elim: wfactorsE)
  | 
| 
 | 
  2278  | 
  from bfs have birr: "\<forall>b \<in> set bs. irreducible G b" by (fast elim: wfactorsE)
  | 
| 
 | 
  2279  | 
  | 
| 
 | 
  2280  | 
  have "\<exists>c cs. c \<in> carrier G \<and> set cs \<subseteq> carrier G \<and> wfactors G cs c \<and> fmset G cs = fmset G bs - fmset G as"
  | 
| 
 | 
  2281  | 
  proof (intro mset_wfactorsEx, simp)
  | 
| 
 | 
  2282  | 
    fix X
  | 
| 
 | 
  2283  | 
    assume "count (fmset G as) X < count (fmset G bs) X"
  | 
| 
 | 
  2284  | 
    hence "0 < count (fmset G bs) X" by simp
  | 
| 
 | 
  2285  | 
    hence "X \<in> set_of (fmset G bs)" by simp
  | 
| 
 | 
  2286  | 
    hence "X \<in> set (map (assocs G) bs)" by (simp add: fmset_def)
  | 
| 
 | 
  2287  | 
    hence "\<exists>x. x \<in> set bs \<and> X = assocs G x" by (induct bs) auto
  | 
| 
 | 
  2288  | 
    from this obtain x
  | 
| 
 | 
  2289  | 
        where xbs: "x \<in> set bs"
  | 
| 
 | 
  2290  | 
        and X: "X = assocs G x"
  | 
| 
 | 
  2291  | 
        by auto
  | 
| 
 | 
  2292  | 
  | 
| 
 | 
  2293  | 
    with bscarr have xcarr: "x \<in> carrier G" by fast
  | 
| 
 | 
  2294  | 
    from xbs birr have xirr: "irreducible G x" by simp
  | 
| 
 | 
  2295  | 
  | 
| 
 | 
  2296  | 
    from xcarr and xirr and X
  | 
| 
 | 
  2297  | 
        show "\<exists>x. x \<in> carrier G \<and> irreducible G x \<and> X = assocs G x" by fast
  | 
| 
 | 
  2298  | 
  qed
  | 
| 
 | 
  2299  | 
  from this obtain c cs
  | 
| 
 | 
  2300  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2301  | 
      and cscarr: "set cs \<subseteq> carrier G" 
  | 
| 
 | 
  2302  | 
      and csf: "wfactors G cs c"
  | 
| 
 | 
  2303  | 
      and csmset: "fmset G cs = fmset G bs - fmset G as" by auto
  | 
| 
 | 
  2304  | 
  | 
| 
 | 
  2305  | 
  from csmset msubset
  | 
| 
 | 
  2306  | 
      have "fmset G bs = fmset G as + fmset G cs"
  | 
| 
 | 
  2307  | 
      by (simp add: multiset_eq_conv_count_eq mset_le_def)
  | 
| 
 | 
  2308  | 
  hence basc: "b \<sim> a \<otimes> c"
  | 
| 
 | 
  2309  | 
      by (rule fmset_wfactors_mult) fact+
  | 
| 
 | 
  2310  | 
  | 
| 
 | 
  2311  | 
  thus ?thesis
  | 
| 
 | 
  2312  | 
  proof (elim associatedE2)
  | 
| 
 | 
  2313  | 
    fix u
  | 
| 
 | 
  2314  | 
    assume "u \<in> Units G"  "b = a \<otimes> c \<otimes> u"
  | 
| 
 | 
  2315  | 
    with acarr ccarr
  | 
| 
 | 
  2316  | 
        show "a divides b" by (fast intro: dividesI[of "c \<otimes> u"] m_assoc)
  | 
| 
 | 
  2317  | 
  qed (simp add: acarr bcarr ccarr)+
  | 
| 
 | 
  2318  | 
qed
  | 
| 
 | 
  2319  | 
  | 
| 
 | 
  2320  | 
lemma (in factorial_monoid) divides_as_fmsubset:
  | 
| 
 | 
  2321  | 
  assumes "wfactors G as a" and "wfactors G bs b"
  | 
| 
 | 
  2322  | 
    and "a \<in> carrier G" and "b \<in> carrier G" 
  | 
| 
 | 
  2323  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2324  | 
  shows "a divides b = (fmset G as \<le># fmset G bs)"
  | 
| 
 | 
  2325  | 
using assms
  | 
| 
 | 
  2326  | 
by (blast intro: divides_fmsubset fmsubset_divides)
  | 
| 
 | 
  2327  | 
  | 
| 
 | 
  2328  | 
  | 
| 
 | 
  2329  | 
text {* Proper factors on multisets *}
 | 
| 
 | 
  2330  | 
  | 
| 
 | 
  2331  | 
lemma (in factorial_monoid) fmset_properfactor:
  | 
| 
 | 
  2332  | 
  assumes asubb: "fmset G as \<le># fmset G bs"
  | 
| 
 | 
  2333  | 
    and anb: "fmset G as \<noteq> fmset G bs"
  | 
| 
 | 
  2334  | 
    and "wfactors G as a" and "wfactors G bs b"
  | 
| 
 | 
  2335  | 
    and "a \<in> carrier G" and "b \<in> carrier G"
  | 
| 
 | 
  2336  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2337  | 
  shows "properfactor G a b"
  | 
| 
 | 
  2338  | 
apply (rule properfactorI)
  | 
| 
 | 
  2339  | 
apply (rule fmsubset_divides[of as bs], fact+)
  | 
| 
 | 
  2340  | 
proof
  | 
| 
 | 
  2341  | 
  assume "b divides a"
  | 
| 
 | 
  2342  | 
  hence "fmset G bs \<le># fmset G as"
  | 
| 
 | 
  2343  | 
      by (rule divides_fmsubset) fact+
  | 
| 
 | 
  2344  | 
  with asubb
  | 
| 
 | 
  2345  | 
      have "fmset G as = fmset G bs" by (simp add: mset_le_antisym)
  | 
| 
 | 
  2346  | 
  with anb
  | 
| 
 | 
  2347  | 
      show "False" ..
  | 
| 
 | 
  2348  | 
qed
  | 
| 
 | 
  2349  | 
  | 
| 
 | 
  2350  | 
lemma (in factorial_monoid) properfactor_fmset:
  | 
| 
 | 
  2351  | 
  assumes pf: "properfactor G a b"
  | 
| 
 | 
  2352  | 
    and "wfactors G as a" and "wfactors G bs b"
  | 
| 
 | 
  2353  | 
    and "a \<in> carrier G" and "b \<in> carrier G"
  | 
| 
 | 
  2354  | 
    and "set as \<subseteq> carrier G" and "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2355  | 
  shows "fmset G as \<le># fmset G bs \<and> fmset G as \<noteq> fmset G bs"
  | 
| 
 | 
  2356  | 
using pf
  | 
| 
 | 
  2357  | 
apply (elim properfactorE)
  | 
| 
 | 
  2358  | 
apply rule
  | 
| 
 | 
  2359  | 
 apply (intro divides_fmsubset, assumption)
  | 
| 
 | 
  2360  | 
  apply (rule assms)+
  | 
| 
 | 
  2361  | 
proof
  | 
| 
 | 
  2362  | 
  assume bna: "\<not> b divides a"
  | 
| 
 | 
  2363  | 
  assume "fmset G as = fmset G bs"
  | 
| 
 | 
  2364  | 
  then have "essentially_equal G as bs" by (rule fmset_ee) fact+
  | 
| 
 | 
  2365  | 
  hence "a \<sim> b" by (rule ee_wfactorsD[of as bs]) fact+
  | 
| 
 | 
  2366  | 
  hence "b divides a" by (elim associatedE)
  | 
| 
 | 
  2367  | 
  with bna
  | 
| 
 | 
  2368  | 
      show "False" ..
  | 
| 
 | 
  2369  | 
qed
  | 
| 
 | 
  2370  | 
  | 
| 
 | 
  2371  | 
  | 
| 
 | 
  2372  | 
subsection {* Irreducible elements are prime *}
 | 
| 
 | 
  2373  | 
  | 
| 
 | 
  2374  | 
lemma (in factorial_monoid) irreducible_is_prime:
  | 
| 
 | 
  2375  | 
  assumes pirr: "irreducible G p"
  | 
| 
 | 
  2376  | 
    and pcarr: "p \<in> carrier G"
  | 
| 
 | 
  2377  | 
  shows "prime G p"
  | 
| 
 | 
  2378  | 
using pirr
  | 
| 
 | 
  2379  | 
proof (elim irreducibleE, intro primeI)
  | 
| 
 | 
  2380  | 
  fix a b
  | 
| 
 | 
  2381  | 
  assume acarr: "a \<in> carrier G"  and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  2382  | 
    and pdvdab: "p divides (a \<otimes> b)"
  | 
| 
 | 
  2383  | 
    and pnunit: "p \<notin> Units G"
  | 
| 
 | 
  2384  | 
  assume irreduc[rule_format]:
  | 
| 
 | 
  2385  | 
         "\<forall>b. b \<in> carrier G \<and> properfactor G b p \<longrightarrow> b \<in> Units G"
  | 
| 
 | 
  2386  | 
  from pdvdab
  | 
| 
 | 
  2387  | 
      have "\<exists>c\<in>carrier G. a \<otimes> b = p \<otimes> c" by (rule dividesD)
  | 
| 
 | 
  2388  | 
  from this obtain c
  | 
| 
 | 
  2389  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2390  | 
      and abpc: "a \<otimes> b = p \<otimes> c"
  | 
| 
 | 
  2391  | 
      by auto
  | 
| 
 | 
  2392  | 
  | 
| 
 | 
  2393  | 
  from acarr have "\<exists>fs. set fs \<subseteq> carrier G \<and> wfactors G fs a" by (rule wfactors_exist)
  | 
| 
 | 
  2394  | 
  from this obtain as where ascarr: "set as \<subseteq> carrier G" and afs: "wfactors G as a" by auto
  | 
| 
 | 
  2395  | 
  | 
| 
 | 
  2396  | 
  from bcarr have "\<exists>fs. set fs \<subseteq> carrier G \<and> wfactors G fs b" by (rule wfactors_exist)
  | 
| 
 | 
  2397  | 
  from this obtain bs where bscarr: "set bs \<subseteq> carrier G" and bfs: "wfactors G bs b" by auto
  | 
| 
 | 
  2398  | 
  | 
| 
 | 
  2399  | 
  from ccarr have "\<exists>fs. set fs \<subseteq> carrier G \<and> wfactors G fs c" by (rule wfactors_exist)
  | 
| 
 | 
  2400  | 
  from this obtain cs where cscarr: "set cs \<subseteq> carrier G" and cfs: "wfactors G cs c" by auto
  | 
| 
 | 
  2401  | 
  | 
| 
 | 
  2402  | 
  note carr[simp] = pcarr acarr bcarr ccarr ascarr bscarr cscarr
  | 
| 
 | 
  2403  | 
  | 
| 
 | 
  2404  | 
  from afs and bfs
  | 
| 
 | 
  2405  | 
      have abfs: "wfactors G (as @ bs) (a \<otimes> b)" by (rule wfactors_mult) fact+
  | 
| 
 | 
  2406  | 
  | 
| 
 | 
  2407  | 
  from pirr cfs
  | 
| 
 | 
  2408  | 
      have pcfs: "wfactors G (p # cs) (p \<otimes> c)" by (rule wfactors_mult_single) fact+
  | 
| 
 | 
  2409  | 
  with abpc
  | 
| 
 | 
  2410  | 
      have abfs': "wfactors G (p # cs) (a \<otimes> b)" by simp
  | 
| 
 | 
  2411  | 
  | 
| 
 | 
  2412  | 
  from abfs' abfs
  | 
| 
 | 
  2413  | 
      have "essentially_equal G (p # cs) (as @ bs)"
  | 
| 
 | 
  2414  | 
      by (rule wfactors_unique) simp+
  | 
| 
 | 
  2415  | 
  | 
| 
 | 
  2416  | 
  hence "\<exists>ds. p # cs <~~> ds \<and> ds [\<sim>] (as @ bs)"
  | 
| 
 | 
  2417  | 
      by (fast elim: essentially_equalE)
  | 
| 
 | 
  2418  | 
  from this obtain ds
  | 
| 
 | 
  2419  | 
      where "p # cs <~~> ds"
  | 
| 
 | 
  2420  | 
      and dsassoc: "ds [\<sim>] (as @ bs)"
  | 
| 
 | 
  2421  | 
      by auto
  | 
| 
 | 
  2422  | 
  | 
| 
 | 
  2423  | 
  then have "p \<in> set ds"
  | 
| 
 | 
  2424  | 
       by (simp add: perm_set_eq[symmetric])
  | 
| 
 | 
  2425  | 
  with dsassoc
  | 
| 
 | 
  2426  | 
       have "\<exists>p'. p' \<in> set (as@bs) \<and> p \<sim> p'"
  | 
| 
 | 
  2427  | 
       unfolding list_all2_conv_all_nth set_conv_nth
  | 
| 
 | 
  2428  | 
       by force
  | 
| 
 | 
  2429  | 
  | 
| 
 | 
  2430  | 
  from this obtain p'
  | 
| 
 | 
  2431  | 
       where "p' \<in> set (as@bs)"
  | 
| 
 | 
  2432  | 
       and pp': "p \<sim> p'"
  | 
| 
 | 
  2433  | 
       by auto
  | 
| 
 | 
  2434  | 
  | 
| 
 | 
  2435  | 
  hence "p' \<in> set as \<or> p' \<in> set bs" by simp
  | 
| 
 | 
  2436  | 
  moreover
  | 
| 
 | 
  2437  | 
  {
 | 
| 
 | 
  2438  | 
    assume p'elem: "p' \<in> set as"
  | 
| 
 | 
  2439  | 
    with ascarr have [simp]: "p' \<in> carrier G" by fast
  | 
| 
 | 
  2440  | 
  | 
| 
 | 
  2441  | 
    note pp'
  | 
| 
 | 
  2442  | 
    also from afs
  | 
| 
 | 
  2443  | 
         have "p' divides a" by (rule wfactors_dividesI) fact+
  | 
| 
 | 
  2444  | 
    finally
  | 
| 
 | 
  2445  | 
         have "p divides a" by simp
  | 
| 
 | 
  2446  | 
  }
  | 
| 
 | 
  2447  | 
  moreover
  | 
| 
 | 
  2448  | 
  {
 | 
| 
 | 
  2449  | 
    assume p'elem: "p' \<in> set bs"
  | 
| 
 | 
  2450  | 
    with bscarr have [simp]: "p' \<in> carrier G" by fast
  | 
| 
 | 
  2451  | 
  | 
| 
 | 
  2452  | 
    note pp'
  | 
| 
 | 
  2453  | 
    also from bfs
  | 
| 
 | 
  2454  | 
         have "p' divides b" by (rule wfactors_dividesI) fact+
  | 
| 
 | 
  2455  | 
    finally
  | 
| 
 | 
  2456  | 
         have "p divides b" by simp
  | 
| 
 | 
  2457  | 
  }
  | 
| 
 | 
  2458  | 
  ultimately
  | 
| 
 | 
  2459  | 
      show "p divides a \<or> p divides b" by fast
  | 
| 
 | 
  2460  | 
qed
  | 
| 
 | 
  2461  | 
  | 
| 
 | 
  2462  | 
  | 
| 
 | 
  2463  | 
--"A version using @{const factors}, more complicated"
 | 
| 
 | 
  2464  | 
lemma (in factorial_monoid) factors_irreducible_is_prime:
  | 
| 
 | 
  2465  | 
  assumes pirr: "irreducible G p"
  | 
| 
 | 
  2466  | 
    and pcarr: "p \<in> carrier G"
  | 
| 
 | 
  2467  | 
  shows "prime G p"
  | 
| 
 | 
  2468  | 
using pirr
  | 
| 
 | 
  2469  | 
apply (elim irreducibleE, intro primeI)
  | 
| 
 | 
  2470  | 
 apply assumption
  | 
| 
 | 
  2471  | 
proof -
  | 
| 
 | 
  2472  | 
  fix a b
  | 
| 
 | 
  2473  | 
  assume acarr: "a \<in> carrier G" 
  | 
| 
 | 
  2474  | 
    and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  2475  | 
    and pdvdab: "p divides (a \<otimes> b)"
  | 
| 
 | 
  2476  | 
  assume irreduc[rule_format]:
  | 
| 
 | 
  2477  | 
         "\<forall>b. b \<in> carrier G \<and> properfactor G b p \<longrightarrow> b \<in> Units G"
  | 
| 
 | 
  2478  | 
  from pdvdab
  | 
| 
 | 
  2479  | 
      have "\<exists>c\<in>carrier G. a \<otimes> b = p \<otimes> c" by (rule dividesD)
  | 
| 
 | 
  2480  | 
  from this obtain c
  | 
| 
 | 
  2481  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2482  | 
      and abpc: "a \<otimes> b = p \<otimes> c"
  | 
| 
 | 
  2483  | 
      by auto
  | 
| 
 | 
  2484  | 
  note [simp] = pcarr acarr bcarr ccarr
  | 
| 
 | 
  2485  | 
  | 
| 
 | 
  2486  | 
  show "p divides a \<or> p divides b"
  | 
| 
 | 
  2487  | 
  proof (cases "a \<in> Units G")
  | 
| 
 | 
  2488  | 
    assume aunit: "a \<in> Units G"
  | 
| 
 | 
  2489  | 
  | 
| 
 | 
  2490  | 
    note pdvdab
  | 
| 
 | 
  2491  | 
    also have "a \<otimes> b = b \<otimes> a" by (simp add: m_comm)
  | 
| 
 | 
  2492  | 
    also from aunit
  | 
| 
 | 
  2493  | 
         have bab: "b \<otimes> a \<sim> b"
  | 
| 
 | 
  2494  | 
         by (intro associatedI2[of "a"], simp+)
  | 
| 
 | 
  2495  | 
    finally
  | 
| 
 | 
  2496  | 
         have "p divides b" by simp
  | 
| 
 | 
  2497  | 
    thus "p divides a \<or> p divides b" ..
  | 
| 
 | 
  2498  | 
  next
  | 
| 
 | 
  2499  | 
    assume anunit: "a \<notin> Units G"
  | 
| 
 | 
  2500  | 
  | 
| 
 | 
  2501  | 
    show "p divides a \<or> p divides b"
  | 
| 
 | 
  2502  | 
    proof (cases "b \<in> Units G")
  | 
| 
 | 
  2503  | 
      assume bunit: "b \<in> Units G"
  | 
| 
 | 
  2504  | 
  | 
| 
 | 
  2505  | 
      note pdvdab
  | 
| 
 | 
  2506  | 
      also from bunit
  | 
| 
 | 
  2507  | 
           have baa: "a \<otimes> b \<sim> a"
  | 
| 
 | 
  2508  | 
           by (intro associatedI2[of "b"], simp+)
  | 
| 
 | 
  2509  | 
      finally
  | 
| 
 | 
  2510  | 
           have "p divides a" by simp
  | 
| 
 | 
  2511  | 
      thus "p divides a \<or> p divides b" ..
  | 
| 
 | 
  2512  | 
    next
  | 
| 
 | 
  2513  | 
      assume bnunit: "b \<notin> Units G"
  | 
| 
 | 
  2514  | 
  | 
| 
 | 
  2515  | 
      have cnunit: "c \<notin> Units G"
  | 
| 
 | 
  2516  | 
      proof (rule ccontr, simp)
  | 
| 
 | 
  2517  | 
        assume cunit: "c \<in> Units G"
  | 
| 
 | 
  2518  | 
        from bnunit
  | 
| 
 | 
  2519  | 
             have "properfactor G a (a \<otimes> b)"
  | 
| 
 | 
  2520  | 
             by (intro properfactorI3[of _ _ b], simp+)
  | 
| 
 | 
  2521  | 
        also note abpc
  | 
| 
 | 
  2522  | 
        also from cunit
  | 
| 
 | 
  2523  | 
             have "p \<otimes> c \<sim> p"
  | 
| 
 | 
  2524  | 
             by (intro associatedI2[of c], simp+)
  | 
| 
 | 
  2525  | 
        finally
  | 
| 
 | 
  2526  | 
             have "properfactor G a p" by simp
  | 
| 
 | 
  2527  | 
  | 
| 
 | 
  2528  | 
        with acarr
  | 
| 
 | 
  2529  | 
             have "a \<in> Units G" by (fast intro: irreduc)
  | 
| 
 | 
  2530  | 
        with anunit
  | 
| 
 | 
  2531  | 
             show "False" ..
  | 
| 
 | 
  2532  | 
      qed
  | 
| 
 | 
  2533  | 
  | 
| 
 | 
  2534  | 
      have abnunit: "a \<otimes> b \<notin> Units G"
  | 
| 
 | 
  2535  | 
      proof clarsimp
  | 
| 
 | 
  2536  | 
        assume abunit: "a \<otimes> b \<in> Units G"
  | 
| 
 | 
  2537  | 
        hence "a \<in> Units G" by (rule unit_factor) fact+
  | 
| 
 | 
  2538  | 
        with anunit
  | 
| 
 | 
  2539  | 
             show "False" ..
  | 
| 
 | 
  2540  | 
      qed
  | 
| 
 | 
  2541  | 
  | 
| 
 | 
  2542  | 
      from acarr anunit have "\<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs a" by (rule factors_exist)
  | 
| 
 | 
  2543  | 
      then obtain as where ascarr: "set as \<subseteq> carrier G" and afac: "factors G as a" by auto
  | 
| 
 | 
  2544  | 
  | 
| 
 | 
  2545  | 
      from bcarr bnunit have "\<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs b" by (rule factors_exist)
  | 
| 
 | 
  2546  | 
      then obtain bs where bscarr: "set bs \<subseteq> carrier G" and bfac: "factors G bs b" by auto
  | 
| 
 | 
  2547  | 
  | 
| 
 | 
  2548  | 
      from ccarr cnunit have "\<exists>fs. set fs \<subseteq> carrier G \<and> factors G fs c" by (rule factors_exist)
  | 
| 
 | 
  2549  | 
      then obtain cs where cscarr: "set cs \<subseteq> carrier G" and cfac: "factors G cs c" by auto
  | 
| 
 | 
  2550  | 
  | 
| 
 | 
  2551  | 
      note [simp] = ascarr bscarr cscarr
  | 
| 
 | 
  2552  | 
  | 
| 
 | 
  2553  | 
      from afac and bfac
  | 
| 
 | 
  2554  | 
          have abfac: "factors G (as @ bs) (a \<otimes> b)" by (rule factors_mult) fact+
  | 
| 
 | 
  2555  | 
  | 
| 
 | 
  2556  | 
      from pirr cfac
  | 
| 
 | 
  2557  | 
          have pcfac: "factors G (p # cs) (p \<otimes> c)" by (rule factors_mult_single) fact+
  | 
| 
 | 
  2558  | 
      with abpc
  | 
| 
 | 
  2559  | 
          have abfac': "factors G (p # cs) (a \<otimes> b)" by simp
  | 
| 
 | 
  2560  | 
  | 
| 
 | 
  2561  | 
      from abfac' abfac
  | 
| 
 | 
  2562  | 
          have "essentially_equal G (p # cs) (as @ bs)"
  | 
| 
 | 
  2563  | 
          by (rule factors_unique) (fact | simp)+
  | 
| 
 | 
  2564  | 
  | 
| 
 | 
  2565  | 
      hence "\<exists>ds. p # cs <~~> ds \<and> ds [\<sim>] (as @ bs)"
  | 
| 
 | 
  2566  | 
          by (fast elim: essentially_equalE)
  | 
| 
 | 
  2567  | 
      from this obtain ds
  | 
| 
 | 
  2568  | 
          where "p # cs <~~> ds"
  | 
| 
 | 
  2569  | 
          and dsassoc: "ds [\<sim>] (as @ bs)"
  | 
| 
 | 
  2570  | 
          by auto
  | 
| 
 | 
  2571  | 
  | 
| 
 | 
  2572  | 
      then have "p \<in> set ds"
  | 
| 
 | 
  2573  | 
           by (simp add: perm_set_eq[symmetric])
  | 
| 
 | 
  2574  | 
      with dsassoc
  | 
| 
 | 
  2575  | 
           have "\<exists>p'. p' \<in> set (as@bs) \<and> p \<sim> p'"
  | 
| 
 | 
  2576  | 
           unfolding list_all2_conv_all_nth set_conv_nth
  | 
| 
 | 
  2577  | 
           by force
  | 
| 
 | 
  2578  | 
  | 
| 
 | 
  2579  | 
      from this obtain p'
  | 
| 
 | 
  2580  | 
	  where "p' \<in> set (as@bs)"
  | 
| 
 | 
  2581  | 
	  and pp': "p \<sim> p'" by auto
  | 
| 
 | 
  2582  | 
  | 
| 
 | 
  2583  | 
      hence "p' \<in> set as \<or> p' \<in> set bs" by simp
  | 
| 
 | 
  2584  | 
      moreover
  | 
| 
 | 
  2585  | 
      {
 | 
| 
 | 
  2586  | 
	assume p'elem: "p' \<in> set as"
  | 
| 
 | 
  2587  | 
	with ascarr have [simp]: "p' \<in> carrier G" by fast
  | 
| 
 | 
  2588  | 
  | 
| 
 | 
  2589  | 
	note pp'
  | 
| 
 | 
  2590  | 
	also from afac p'elem
  | 
| 
 | 
  2591  | 
	     have "p' divides a" by (rule factors_dividesI) fact+
  | 
| 
 | 
  2592  | 
	finally
  | 
| 
 | 
  2593  | 
	     have "p divides a" by simp
  | 
| 
 | 
  2594  | 
      }
  | 
| 
 | 
  2595  | 
      moreover
  | 
| 
 | 
  2596  | 
      {
 | 
| 
 | 
  2597  | 
	assume p'elem: "p' \<in> set bs"
  | 
| 
 | 
  2598  | 
	with bscarr have [simp]: "p' \<in> carrier G" by fast
  | 
| 
 | 
  2599  | 
  | 
| 
 | 
  2600  | 
	note pp'
  | 
| 
 | 
  2601  | 
	also from bfac
  | 
| 
 | 
  2602  | 
	     have "p' divides b" by (rule factors_dividesI) fact+
  | 
| 
 | 
  2603  | 
	finally have "p divides b" by simp
  | 
| 
 | 
  2604  | 
      }
  | 
| 
 | 
  2605  | 
      ultimately
  | 
| 
 | 
  2606  | 
	  show "p divides a \<or> p divides b" by fast
  | 
| 
 | 
  2607  | 
    qed
  | 
| 
 | 
  2608  | 
  qed
  | 
| 
 | 
  2609  | 
qed
  | 
| 
 | 
  2610  | 
  | 
| 
 | 
  2611  | 
  | 
| 
 | 
  2612  | 
subsection {* Greatest common divisors and lowest common multiples *}
 | 
| 
 | 
  2613  | 
  | 
| 
 | 
  2614  | 
subsubsection {* Definitions *}
 | 
| 
 | 
  2615  | 
  | 
| 
 | 
  2616  | 
constdefs (structure G)
  | 
| 
 | 
  2617  | 
  isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \<Rightarrow> bool"  ("(_ gcdof\<index> _ _)" [81,81,81] 80)
 | 
| 
 | 
  2618  | 
  "x gcdof a b \<equiv> x divides a \<and> x divides b \<and> 
  | 
| 
 | 
  2619  | 
                 (\<forall>y\<in>carrier G. (y divides a \<and> y divides b \<longrightarrow> y divides x))"
  | 
| 
 | 
  2620  | 
  | 
| 
 | 
  2621  | 
  islcm :: "[_, 'a, 'a, 'a] \<Rightarrow> bool"  ("(_ lcmof\<index> _ _)" [81,81,81] 80)
 | 
| 
 | 
  2622  | 
  "x lcmof a b \<equiv> a divides x \<and> b divides x \<and> 
  | 
| 
 | 
  2623  | 
                 (\<forall>y\<in>carrier G. (a divides y \<and> b divides y \<longrightarrow> x divides y))"
  | 
| 
 | 
  2624  | 
  | 
| 
 | 
  2625  | 
constdefs (structure G)
  | 
| 
 | 
  2626  | 
  somegcd :: "('a,_) monoid_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 
 | 
  2627  | 
  "somegcd G a b == SOME x. x \<in> carrier G \<and> x gcdof a b"
  | 
| 
 | 
  2628  | 
  | 
| 
 | 
  2629  | 
  somelcm :: "('a,_) monoid_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
 | 
| 
 | 
  2630  | 
  "somelcm G a b == SOME x. x \<in> carrier G \<and> x lcmof a b"
  | 
| 
 | 
  2631  | 
  | 
| 
 | 
  2632  | 
constdefs (structure G)
  | 
| 
 | 
  2633  | 
  "SomeGcd G A == ginf (division_rel G) A"
  | 
| 
 | 
  2634  | 
  | 
| 
 | 
  2635  | 
  | 
| 
 | 
  2636  | 
locale gcd_condition_monoid = comm_monoid_cancel +
  | 
| 
 | 
  2637  | 
  assumes gcdof_exists:
  | 
| 
 | 
  2638  | 
          "\<lbrakk>a \<in> carrier G; b \<in> carrier G\<rbrakk> \<Longrightarrow> \<exists>c. c \<in> carrier G \<and> c gcdof a b"
  | 
| 
 | 
  2639  | 
  | 
| 
 | 
  2640  | 
locale primeness_condition_monoid = comm_monoid_cancel +
  | 
| 
 | 
  2641  | 
  assumes irreducible_prime:
  | 
| 
 | 
  2642  | 
          "\<lbrakk>a \<in> carrier G; irreducible G a\<rbrakk> \<Longrightarrow> prime G a"
  | 
| 
 | 
  2643  | 
  | 
| 
 | 
  2644  | 
locale divisor_chain_condition_monoid = comm_monoid_cancel +
  | 
| 
 | 
  2645  | 
  assumes division_wellfounded:
  | 
| 
 | 
  2646  | 
          "wf {(x, y). x \<in> carrier G \<and> y \<in> carrier G \<and> properfactor G x y}"
 | 
| 
 | 
  2647  | 
  | 
| 
 | 
  2648  | 
  | 
| 
 | 
  2649  | 
subsubsection {* Connections to \texttt{Lattice.thy} *}
 | 
| 
 | 
  2650  | 
  | 
| 
 | 
  2651  | 
lemma gcdof_ggreatestLower:
  | 
| 
 | 
  2652  | 
  fixes G (structure)
  | 
| 
 | 
  2653  | 
  assumes carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2654  | 
  shows "(x \<in> carrier G \<and> x gcdof a b) =
  | 
| 
 | 
  2655  | 
         ggreatest (division_rel G) x (Lower (division_rel G) {a, b})"
 | 
| 
 | 
  2656  | 
unfolding isgcd_def ggreatest_def Lower_def elem_def
  | 
| 
 | 
  2657  | 
proof (simp, safe)
  | 
| 
 | 
  2658  | 
  fix xa
  | 
| 
 | 
  2659  | 
  assume r1[rule_format]: "\<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> xa divides x"
  | 
| 
 | 
  2660  | 
  assume r2[rule_format]: "\<forall>y\<in>carrier G. y divides a \<and> y divides b \<longrightarrow> y divides x"
  | 
| 
 | 
  2661  | 
  | 
| 
 | 
  2662  | 
  assume "xa \<in> carrier G"  "x divides a"  "x divides b"
  | 
| 
 | 
  2663  | 
  with carr
  | 
| 
 | 
  2664  | 
  show "xa divides x"
  | 
| 
 | 
  2665  | 
      by (fast intro: r1 r2)
  | 
| 
 | 
  2666  | 
next
  | 
| 
 | 
  2667  | 
  fix a' y
  | 
| 
 | 
  2668  | 
  assume r1[rule_format]:
  | 
| 
 | 
  2669  | 
         "\<forall>xa\<in>{l. \<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> l divides x} \<inter> carrier G.
 | 
| 
 | 
  2670  | 
           xa divides x"
  | 
| 
 | 
  2671  | 
  assume "y \<in> carrier G"  "y divides a"  "y divides b"
  | 
| 
 | 
  2672  | 
  with carr
  | 
| 
 | 
  2673  | 
       show "y divides x"
  | 
| 
 | 
  2674  | 
       by (fast intro: r1)
  | 
| 
 | 
  2675  | 
qed (simp, simp)
  | 
| 
 | 
  2676  | 
  | 
| 
 | 
  2677  | 
lemma lcmof_gleastUpper:
  | 
| 
 | 
  2678  | 
  fixes G (structure)
  | 
| 
 | 
  2679  | 
  assumes carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2680  | 
  shows "(x \<in> carrier G \<and> x lcmof a b) =
  | 
| 
 | 
  2681  | 
         gleast (division_rel G) x (Upper (division_rel G) {a, b})"
 | 
| 
 | 
  2682  | 
unfolding islcm_def gleast_def Upper_def elem_def
  | 
| 
 | 
  2683  | 
proof (simp, safe)
  | 
| 
 | 
  2684  | 
  fix xa
  | 
| 
 | 
  2685  | 
  assume r1[rule_format]: "\<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> x divides xa"
  | 
| 
 | 
  2686  | 
  assume r2[rule_format]: "\<forall>y\<in>carrier G. a divides y \<and> b divides y \<longrightarrow> x divides y"
  | 
| 
 | 
  2687  | 
  | 
| 
 | 
  2688  | 
  assume "xa \<in> carrier G"  "a divides x"  "b divides x"
  | 
| 
 | 
  2689  | 
  with carr
  | 
| 
 | 
  2690  | 
  show "x divides xa"
  | 
| 
 | 
  2691  | 
      by (fast intro: r1 r2)
  | 
| 
 | 
  2692  | 
next
  | 
| 
 | 
  2693  | 
  fix a' y
  | 
| 
 | 
  2694  | 
  assume r1[rule_format]:
  | 
| 
 | 
  2695  | 
         "\<forall>xa\<in>{l. \<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> x divides l} \<inter> carrier G.
 | 
| 
 | 
  2696  | 
           x divides xa"
  | 
| 
 | 
  2697  | 
  assume "y \<in> carrier G"  "a divides y"  "b divides y"
  | 
| 
 | 
  2698  | 
  with carr
  | 
| 
 | 
  2699  | 
       show "x divides y"
  | 
| 
 | 
  2700  | 
       by (fast intro: r1)
  | 
| 
 | 
  2701  | 
qed (simp, simp)
  | 
| 
 | 
  2702  | 
  | 
| 
 | 
  2703  | 
lemma somegcd_gmeet:
  | 
| 
 | 
  2704  | 
  fixes G (structure)
  | 
| 
 | 
  2705  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2706  | 
  shows "somegcd G a b = gmeet (division_rel G) a b"
  | 
| 
 | 
  2707  | 
unfolding somegcd_def gmeet_def ginf_def
  | 
| 
 | 
  2708  | 
by (simp add: gcdof_ggreatestLower[OF carr])
  | 
| 
 | 
  2709  | 
  | 
| 
 | 
  2710  | 
lemma (in monoid) isgcd_divides_l:
  | 
| 
 | 
  2711  | 
  assumes "a divides b"
  | 
| 
 | 
  2712  | 
    and "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2713  | 
  shows "a gcdof a b"
  | 
| 
 | 
  2714  | 
using assms
  | 
| 
 | 
  2715  | 
unfolding isgcd_def
  | 
| 
 | 
  2716  | 
by fast
  | 
| 
 | 
  2717  | 
  | 
| 
 | 
  2718  | 
lemma (in monoid) isgcd_divides_r:
  | 
| 
 | 
  2719  | 
  assumes "b divides a"
  | 
| 
 | 
  2720  | 
    and "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2721  | 
  shows "b gcdof a b"
  | 
| 
 | 
  2722  | 
using assms
  | 
| 
 | 
  2723  | 
unfolding isgcd_def
  | 
| 
 | 
  2724  | 
by fast
  | 
| 
 | 
  2725  | 
  | 
| 
 | 
  2726  | 
  | 
| 
 | 
  2727  | 
subsubsection {* Existence of gcd and lcm *}
 | 
| 
 | 
  2728  | 
  | 
| 
 | 
  2729  | 
lemma (in factorial_monoid) gcdof_exists:
  | 
| 
 | 
  2730  | 
  assumes acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  2731  | 
  shows "\<exists>c. c \<in> carrier G \<and> c gcdof a b"
  | 
| 
 | 
  2732  | 
proof -
  | 
| 
 | 
  2733  | 
  from acarr have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a" by (rule wfactors_exist)
  | 
| 
 | 
  2734  | 
  from this obtain as
  | 
| 
 | 
  2735  | 
      where ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  2736  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  2737  | 
      by auto
  | 
| 
 | 
  2738  | 
  from afs have airr: "\<forall>a \<in> set as. irreducible G a" by (fast elim: wfactorsE)
  | 
| 
 | 
  2739  | 
  | 
| 
 | 
  2740  | 
  from bcarr have "\<exists>bs. set bs \<subseteq> carrier G \<and> wfactors G bs b" by (rule wfactors_exist)
  | 
| 
 | 
  2741  | 
  from this obtain bs
  | 
| 
 | 
  2742  | 
      where bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2743  | 
      and bfs: "wfactors G bs b"
  | 
| 
 | 
  2744  | 
      by auto
  | 
| 
 | 
  2745  | 
  from bfs have birr: "\<forall>b \<in> set bs. irreducible G b" by (fast elim: wfactorsE)
  | 
| 
 | 
  2746  | 
  | 
| 
 | 
  2747  | 
  have "\<exists>c cs. c \<in> carrier G \<and> set cs \<subseteq> carrier G \<and> wfactors G cs c \<and> 
  | 
| 
 | 
  2748  | 
               fmset G cs = fmset G as #\<inter> fmset G bs"
  | 
| 
 | 
  2749  | 
  proof (intro mset_wfactorsEx)
  | 
| 
 | 
  2750  | 
    fix X
  | 
| 
 | 
  2751  | 
    assume "X \<in> set_of (fmset G as #\<inter> fmset G bs)"
  | 
| 
 | 
  2752  | 
    hence "X \<in> set_of (fmset G as)" by (simp add: multiset_inter_def)
  | 
| 
 | 
  2753  | 
    hence "X \<in> set (map (assocs G) as)" by (simp add: fmset_def)
  | 
| 
 | 
  2754  | 
    hence "\<exists>x. X = assocs G x \<and> x \<in> set as" by (induct as) auto
  | 
| 
 | 
  2755  | 
    from this obtain x
  | 
| 
 | 
  2756  | 
        where X: "X = assocs G x"
  | 
| 
 | 
  2757  | 
        and xas: "x \<in> set as"
  | 
| 
 | 
  2758  | 
        by auto
  | 
| 
 | 
  2759  | 
    with ascarr have xcarr: "x \<in> carrier G" by fast
  | 
| 
 | 
  2760  | 
    from xas airr have xirr: "irreducible G x" by simp
  | 
| 
 | 
  2761  | 
 
  | 
| 
 | 
  2762  | 
    from xcarr and xirr and X
  | 
| 
 | 
  2763  | 
        show "\<exists>x. (x \<in> carrier G \<and> irreducible G x) \<and> X = assocs G x" by fast
  | 
| 
 | 
  2764  | 
  qed
  | 
| 
 | 
  2765  | 
  | 
| 
 | 
  2766  | 
  from this obtain c cs
  | 
| 
 | 
  2767  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2768  | 
      and cscarr: "set cs \<subseteq> carrier G" 
  | 
| 
 | 
  2769  | 
      and csirr: "wfactors G cs c"
  | 
| 
 | 
  2770  | 
      and csmset: "fmset G cs = fmset G as #\<inter> fmset G bs" by auto
  | 
| 
 | 
  2771  | 
  | 
| 
 | 
  2772  | 
  have "c gcdof a b"
  | 
| 
 | 
  2773  | 
  proof (simp add: isgcd_def, safe)
  | 
| 
 | 
  2774  | 
    from csmset
  | 
| 
 | 
  2775  | 
        have "fmset G cs \<le># fmset G as"
  | 
| 
 | 
  2776  | 
        by (simp add: multiset_inter_def mset_le_def)
  | 
| 
 | 
  2777  | 
    thus "c divides a" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2778  | 
  next
  | 
| 
 | 
  2779  | 
    from csmset
  | 
| 
 | 
  2780  | 
        have "fmset G cs \<le># fmset G bs"
  | 
| 
 | 
  2781  | 
        by (simp add: multiset_inter_def mset_le_def, force)
  | 
| 
 | 
  2782  | 
    thus "c divides b" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2783  | 
  next
  | 
| 
 | 
  2784  | 
    fix y
  | 
| 
 | 
  2785  | 
    assume ycarr: "y \<in> carrier G"
  | 
| 
 | 
  2786  | 
    hence "\<exists>ys. set ys \<subseteq> carrier G \<and> wfactors G ys y" by (rule wfactors_exist)
  | 
| 
 | 
  2787  | 
    from this obtain ys
  | 
| 
 | 
  2788  | 
        where yscarr: "set ys \<subseteq> carrier G"
  | 
| 
 | 
  2789  | 
        and yfs: "wfactors G ys y"
  | 
| 
 | 
  2790  | 
        by auto
  | 
| 
 | 
  2791  | 
  | 
| 
 | 
  2792  | 
    assume "y divides a"
  | 
| 
 | 
  2793  | 
    hence ya: "fmset G ys \<le># fmset G as" by (rule divides_fmsubset) fact+
  | 
| 
 | 
  2794  | 
  | 
| 
 | 
  2795  | 
    assume "y divides b"
  | 
| 
 | 
  2796  | 
    hence yb: "fmset G ys \<le># fmset G bs" by (rule divides_fmsubset) fact+
  | 
| 
 | 
  2797  | 
  | 
| 
 | 
  2798  | 
    from ya yb csmset
  | 
| 
 | 
  2799  | 
    have "fmset G ys \<le># fmset G cs" by (simp add: mset_le_def multiset_inter_count)
  | 
| 
 | 
  2800  | 
    thus "y divides c" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2801  | 
  qed
  | 
| 
 | 
  2802  | 
  | 
| 
 | 
  2803  | 
  with ccarr
  | 
| 
 | 
  2804  | 
      show "\<exists>c. c \<in> carrier G \<and> c gcdof a b" by fast
  | 
| 
 | 
  2805  | 
qed
  | 
| 
 | 
  2806  | 
  | 
| 
 | 
  2807  | 
lemma (in factorial_monoid) lcmof_exists:
  | 
| 
 | 
  2808  | 
  assumes acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  2809  | 
  shows "\<exists>c. c \<in> carrier G \<and> c lcmof a b"
  | 
| 
 | 
  2810  | 
proof -
  | 
| 
 | 
  2811  | 
  from acarr have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a" by (rule wfactors_exist)
  | 
| 
 | 
  2812  | 
  from this obtain as
  | 
| 
 | 
  2813  | 
      where ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  2814  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  2815  | 
      by auto
  | 
| 
 | 
  2816  | 
  from afs have airr: "\<forall>a \<in> set as. irreducible G a" by (fast elim: wfactorsE)
  | 
| 
 | 
  2817  | 
  | 
| 
 | 
  2818  | 
  from bcarr have "\<exists>bs. set bs \<subseteq> carrier G \<and> wfactors G bs b" by (rule wfactors_exist)
  | 
| 
 | 
  2819  | 
  from this obtain bs
  | 
| 
 | 
  2820  | 
      where bscarr: "set bs \<subseteq> carrier G"
  | 
| 
 | 
  2821  | 
      and bfs: "wfactors G bs b"
  | 
| 
 | 
  2822  | 
      by auto
  | 
| 
 | 
  2823  | 
  from bfs have birr: "\<forall>b \<in> set bs. irreducible G b" by (fast elim: wfactorsE)
  | 
| 
 | 
  2824  | 
  | 
| 
 | 
  2825  | 
  have "\<exists>c cs. c \<in> carrier G \<and> set cs \<subseteq> carrier G \<and> wfactors G cs c \<and> 
  | 
| 
 | 
  2826  | 
               fmset G cs = (fmset G as - fmset G bs) + fmset G bs"
  | 
| 
 | 
  2827  | 
  proof (intro mset_wfactorsEx)
  | 
| 
 | 
  2828  | 
    fix X
  | 
| 
 | 
  2829  | 
    assume "X \<in> set_of ((fmset G as - fmset G bs) + fmset G bs)"
  | 
| 
 | 
  2830  | 
    hence "X \<in> set_of (fmset G as) \<or> X \<in> set_of (fmset G bs)"
  | 
| 
 | 
  2831  | 
       by (cases "X :# fmset G bs", simp, simp)
  | 
| 
 | 
  2832  | 
    moreover
  | 
| 
 | 
  2833  | 
    {
 | 
| 
 | 
  2834  | 
      assume "X \<in> set_of (fmset G as)"
  | 
| 
 | 
  2835  | 
      hence "X \<in> set (map (assocs G) as)" by (simp add: fmset_def)
  | 
| 
 | 
  2836  | 
      hence "\<exists>x. x \<in> set as \<and> X = assocs G x" by (induct as) auto
  | 
| 
 | 
  2837  | 
      from this obtain x
  | 
| 
 | 
  2838  | 
          where xas: "x \<in> set as"
  | 
| 
 | 
  2839  | 
          and X: "X = assocs G x" by auto
  | 
| 
 | 
  2840  | 
  | 
| 
 | 
  2841  | 
      with ascarr have xcarr: "x \<in> carrier G" by fast
  | 
| 
 | 
  2842  | 
      from xas airr have xirr: "irreducible G x" by simp
  | 
| 
 | 
  2843  | 
  | 
| 
 | 
  2844  | 
      from xcarr and xirr and X
  | 
| 
 | 
  2845  | 
          have "\<exists>x. (x \<in> carrier G \<and> irreducible G x) \<and> X = assocs G x" by fast
  | 
| 
 | 
  2846  | 
    }
  | 
| 
 | 
  2847  | 
    moreover
  | 
| 
 | 
  2848  | 
    {
 | 
| 
 | 
  2849  | 
      assume "X \<in> set_of (fmset G bs)"
  | 
| 
 | 
  2850  | 
      hence "X \<in> set (map (assocs G) bs)" by (simp add: fmset_def)
  | 
| 
 | 
  2851  | 
      hence "\<exists>x. x \<in> set bs \<and> X = assocs G x" by (induct as) auto
  | 
| 
 | 
  2852  | 
      from this obtain x
  | 
| 
 | 
  2853  | 
          where xbs: "x \<in> set bs"
  | 
| 
 | 
  2854  | 
          and X: "X = assocs G x" by auto
  | 
| 
 | 
  2855  | 
  | 
| 
 | 
  2856  | 
      with bscarr have xcarr: "x \<in> carrier G" by fast
  | 
| 
 | 
  2857  | 
      from xbs birr have xirr: "irreducible G x" by simp
  | 
| 
 | 
  2858  | 
  | 
| 
 | 
  2859  | 
      from xcarr and xirr and X
  | 
| 
 | 
  2860  | 
          have "\<exists>x. (x \<in> carrier G \<and> irreducible G x) \<and> X = assocs G x" by fast
  | 
| 
 | 
  2861  | 
    }
  | 
| 
 | 
  2862  | 
    ultimately
  | 
| 
 | 
  2863  | 
    show "\<exists>x. (x \<in> carrier G \<and> irreducible G x) \<and> X = assocs G x" by fast
  | 
| 
 | 
  2864  | 
  qed
  | 
| 
 | 
  2865  | 
  | 
| 
 | 
  2866  | 
  from this obtain c cs
  | 
| 
 | 
  2867  | 
      where ccarr: "c \<in> carrier G"
  | 
| 
 | 
  2868  | 
      and cscarr: "set cs \<subseteq> carrier G" 
  | 
| 
 | 
  2869  | 
      and csirr: "wfactors G cs c"
  | 
| 
 | 
  2870  | 
      and csmset: "fmset G cs = fmset G as - fmset G bs + fmset G bs" by auto
  | 
| 
 | 
  2871  | 
  | 
| 
 | 
  2872  | 
  have "c lcmof a b"
  | 
| 
 | 
  2873  | 
  proof (simp add: islcm_def, safe)
  | 
| 
 | 
  2874  | 
    from csmset have "fmset G as \<le># fmset G cs" by (simp add: mset_le_def, force)
  | 
| 
 | 
  2875  | 
    thus "a divides c" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2876  | 
  next
  | 
| 
 | 
  2877  | 
    from csmset have "fmset G bs \<le># fmset G cs" by (simp add: mset_le_def)
  | 
| 
 | 
  2878  | 
    thus "b divides c" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2879  | 
  next
  | 
| 
 | 
  2880  | 
    fix y
  | 
| 
 | 
  2881  | 
    assume ycarr: "y \<in> carrier G"
  | 
| 
 | 
  2882  | 
    hence "\<exists>ys. set ys \<subseteq> carrier G \<and> wfactors G ys y" by (rule wfactors_exist)
  | 
| 
 | 
  2883  | 
    from this obtain ys
  | 
| 
 | 
  2884  | 
        where yscarr: "set ys \<subseteq> carrier G"
  | 
| 
 | 
  2885  | 
        and yfs: "wfactors G ys y"
  | 
| 
 | 
  2886  | 
        by auto
  | 
| 
 | 
  2887  | 
  | 
| 
 | 
  2888  | 
    assume "a divides y"
  | 
| 
 | 
  2889  | 
    hence ya: "fmset G as \<le># fmset G ys" by (rule divides_fmsubset) fact+
  | 
| 
 | 
  2890  | 
  | 
| 
 | 
  2891  | 
    assume "b divides y"
  | 
| 
 | 
  2892  | 
    hence yb: "fmset G bs \<le># fmset G ys" by (rule divides_fmsubset) fact+
  | 
| 
 | 
  2893  | 
  | 
| 
 | 
  2894  | 
    from ya yb csmset
  | 
| 
 | 
  2895  | 
    have "fmset G cs \<le># fmset G ys"
  | 
| 
 | 
  2896  | 
      apply (simp add: mset_le_def, clarify)
  | 
| 
 | 
  2897  | 
      apply (case_tac "count (fmset G as) a < count (fmset G bs) a")
  | 
| 
 | 
  2898  | 
       apply simp
  | 
| 
 | 
  2899  | 
      apply simp
  | 
| 
 | 
  2900  | 
    done
  | 
| 
 | 
  2901  | 
    thus "c divides y" by (rule fmsubset_divides) fact+
  | 
| 
 | 
  2902  | 
  qed
  | 
| 
 | 
  2903  | 
  | 
| 
 | 
  2904  | 
  with ccarr
  | 
| 
 | 
  2905  | 
      show "\<exists>c. c \<in> carrier G \<and> c lcmof a b" by fast
  | 
| 
 | 
  2906  | 
qed
  | 
| 
 | 
  2907  | 
  | 
| 
 | 
  2908  | 
  | 
| 
 | 
  2909  | 
subsection {* Conditions for factoriality *}
 | 
| 
 | 
  2910  | 
  | 
| 
 | 
  2911  | 
subsubsection {* Gcd condition *}
 | 
| 
 | 
  2912  | 
  | 
| 
 | 
  2913  | 
lemma (in gcd_condition_monoid) division_glower_semilattice [simp]:
  | 
| 
 | 
  2914  | 
  shows "glower_semilattice (division_rel G)"
  | 
| 
 | 
  2915  | 
proof -
  | 
| 
 | 
  2916  | 
  interpret gpartial_order ["division_rel G"] ..
  | 
| 
 | 
  2917  | 
  show ?thesis
  | 
| 
 | 
  2918  | 
  apply (unfold_locales, simp_all)
  | 
| 
 | 
  2919  | 
  proof -
  | 
| 
 | 
  2920  | 
    fix x y
  | 
| 
 | 
  2921  | 
    assume carr: "x \<in> carrier G"  "y \<in> carrier G"
  | 
| 
 | 
  2922  | 
    hence "\<exists>z. z \<in> carrier G \<and> z gcdof x y" by (rule gcdof_exists)
  | 
| 
 | 
  2923  | 
    from this obtain z
  | 
| 
 | 
  2924  | 
        where zcarr: "z \<in> carrier G"
  | 
| 
 | 
  2925  | 
        and isgcd: "z gcdof x y"
  | 
| 
 | 
  2926  | 
        by auto
  | 
| 
 | 
  2927  | 
    with carr
  | 
| 
 | 
  2928  | 
    have "ggreatest (division_rel G) z (Lower (division_rel G) {x, y})"
 | 
| 
 | 
  2929  | 
        by (subst gcdof_ggreatestLower[symmetric], simp+)
  | 
| 
 | 
  2930  | 
    thus "\<exists>z. ggreatest (division_rel G) z (Lower (division_rel G) {x, y})" by fast
 | 
| 
 | 
  2931  | 
  qed
  | 
| 
 | 
  2932  | 
qed
  | 
| 
 | 
  2933  | 
  | 
| 
 | 
  2934  | 
lemma (in gcd_condition_monoid) gcdof_cong_l:
  | 
| 
 | 
  2935  | 
  assumes a'a: "a' \<sim> a"
  | 
| 
 | 
  2936  | 
    and agcd: "a gcdof b c"
  | 
| 
 | 
  2937  | 
    and a'carr: "a' \<in> carrier G" and carr': "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
  2938  | 
  shows "a' gcdof b c"
  | 
| 
 | 
  2939  | 
proof -
  | 
| 
 | 
  2940  | 
  note carr = a'carr carr'
  | 
| 
 | 
  2941  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  2942  | 
  have "a' \<in> carrier G \<and> a' gcdof b c"
  | 
| 
 | 
  2943  | 
    apply (simp add: gcdof_ggreatestLower carr')
  | 
| 
 | 
  2944  | 
    apply (subst ggreatest_Lower_cong_l[of _ a])
  | 
| 
 | 
  2945  | 
       apply (simp add: a'a)
  | 
| 
 | 
  2946  | 
      apply (simp add: carr)
  | 
| 
 | 
  2947  | 
     apply (simp add: carr)
  | 
| 
 | 
  2948  | 
    apply (simp add: carr)
  | 
| 
 | 
  2949  | 
    apply (simp add: gcdof_ggreatestLower[symmetric] agcd carr)
  | 
| 
 | 
  2950  | 
  done
  | 
| 
 | 
  2951  | 
  thus ?thesis ..
  | 
| 
 | 
  2952  | 
qed
  | 
| 
 | 
  2953  | 
  | 
| 
 | 
  2954  | 
lemma (in gcd_condition_monoid) gcd_closed [simp]:
  | 
| 
 | 
  2955  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2956  | 
  shows "somegcd G a b \<in> carrier G"
  | 
| 
 | 
  2957  | 
proof -
  | 
| 
 | 
  2958  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  2959  | 
  show ?thesis
  | 
| 
 | 
  2960  | 
    apply (simp add: somegcd_gmeet[OF carr])
  | 
| 
 | 
  2961  | 
    apply (rule gmeet_closed[simplified], fact+)
  | 
| 
 | 
  2962  | 
  done
  | 
| 
 | 
  2963  | 
qed
  | 
| 
 | 
  2964  | 
  | 
| 
 | 
  2965  | 
lemma (in gcd_condition_monoid) gcd_isgcd:
  | 
| 
 | 
  2966  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2967  | 
  shows "(somegcd G a b) gcdof a b"
  | 
| 
 | 
  2968  | 
proof -
  | 
| 
 | 
  2969  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  2970  | 
  from carr
  | 
| 
 | 
  2971  | 
  have "somegcd G a b \<in> carrier G \<and> (somegcd G a b) gcdof a b"
  | 
| 
 | 
  2972  | 
    apply (subst gcdof_ggreatestLower, simp, simp)
  | 
| 
 | 
  2973  | 
    apply (simp add: somegcd_gmeet[OF carr] gmeet_def)
  | 
| 
 | 
  2974  | 
    apply (rule ginf_of_two_ggreatest[simplified], assumption+)
  | 
| 
 | 
  2975  | 
  done
  | 
| 
 | 
  2976  | 
  thus "(somegcd G a b) gcdof a b" by simp
  | 
| 
 | 
  2977  | 
qed
  | 
| 
 | 
  2978  | 
  | 
| 
 | 
  2979  | 
lemma (in gcd_condition_monoid) gcd_exists:
  | 
| 
 | 
  2980  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2981  | 
  shows "\<exists>x\<in>carrier G. x = somegcd G a b"
  | 
| 
 | 
  2982  | 
proof -
  | 
| 
 | 
  2983  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  2984  | 
  show ?thesis
  | 
| 
 | 
  2985  | 
    apply (simp add: somegcd_gmeet[OF carr])
  | 
| 
 | 
  2986  | 
    apply (rule gmeet_closed[simplified], fact+)
  | 
| 
 | 
  2987  | 
  done
  | 
| 
 | 
  2988  | 
qed
  | 
| 
 | 
  2989  | 
  | 
| 
 | 
  2990  | 
lemma (in gcd_condition_monoid) gcd_divides_l:
  | 
| 
 | 
  2991  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  2992  | 
  shows "(somegcd G a b) divides a"
  | 
| 
 | 
  2993  | 
proof -
  | 
| 
 | 
  2994  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  2995  | 
  show ?thesis
  | 
| 
 | 
  2996  | 
    apply (simp add: somegcd_gmeet[OF carr])
  | 
| 
 | 
  2997  | 
    apply (rule gmeet_left[simplified], fact+)
  | 
| 
 | 
  2998  | 
  done
  | 
| 
 | 
  2999  | 
qed
  | 
| 
 | 
  3000  | 
  | 
| 
 | 
  3001  | 
lemma (in gcd_condition_monoid) gcd_divides_r:
  | 
| 
 | 
  3002  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  3003  | 
  shows "(somegcd G a b) divides b"
  | 
| 
 | 
  3004  | 
proof -
  | 
| 
 | 
  3005  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3006  | 
  show ?thesis
  | 
| 
 | 
  3007  | 
    apply (simp add: somegcd_gmeet[OF carr])
  | 
| 
 | 
  3008  | 
    apply (rule gmeet_right[simplified], fact+)
  | 
| 
 | 
  3009  | 
  done
  | 
| 
 | 
  3010  | 
qed
  | 
| 
 | 
  3011  | 
  | 
| 
 | 
  3012  | 
lemma (in gcd_condition_monoid) gcd_divides:
  | 
| 
 | 
  3013  | 
  assumes sub: "z divides x"  "z divides y"
  | 
| 
 | 
  3014  | 
    and L: "x \<in> carrier G"  "y \<in> carrier G"  "z \<in> carrier G"
  | 
| 
 | 
  3015  | 
  shows "z divides (somegcd G x y)"
  | 
| 
 | 
  3016  | 
proof -
  | 
| 
 | 
  3017  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3018  | 
  show ?thesis
  | 
| 
 | 
  3019  | 
    apply (simp add: somegcd_gmeet L)
  | 
| 
 | 
  3020  | 
    apply (rule gmeet_le[simplified], fact+)
  | 
| 
 | 
  3021  | 
  done
  | 
| 
 | 
  3022  | 
qed
  | 
| 
 | 
  3023  | 
  | 
| 
 | 
  3024  | 
lemma (in gcd_condition_monoid) gcd_cong_l:
  | 
| 
 | 
  3025  | 
  assumes xx': "x \<sim> x'"
  | 
| 
 | 
  3026  | 
    and carr: "x \<in> carrier G"  "x' \<in> carrier G"  "y \<in> carrier G"
  | 
| 
 | 
  3027  | 
  shows "somegcd G x y \<sim> somegcd G x' y"
  | 
| 
 | 
  3028  | 
proof -
  | 
| 
 | 
  3029  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3030  | 
  show ?thesis
  | 
| 
 | 
  3031  | 
    apply (simp add: somegcd_gmeet carr)
  | 
| 
 | 
  3032  | 
    apply (rule gmeet_cong_l[simplified], fact+)
  | 
| 
 | 
  3033  | 
  done
  | 
| 
 | 
  3034  | 
qed
  | 
| 
 | 
  3035  | 
  | 
| 
 | 
  3036  | 
lemma (in gcd_condition_monoid) gcd_cong_r:
  | 
| 
 | 
  3037  | 
  assumes carr: "x \<in> carrier G"  "y \<in> carrier G"  "y' \<in> carrier G"
  | 
| 
 | 
  3038  | 
    and yy': "y \<sim> y'"
  | 
| 
 | 
  3039  | 
  shows "somegcd G x y \<sim> somegcd G x y'"
  | 
| 
 | 
  3040  | 
proof -
  | 
| 
 | 
  3041  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3042  | 
  show ?thesis
  | 
| 
 | 
  3043  | 
    apply (simp add: somegcd_gmeet carr)
  | 
| 
 | 
  3044  | 
    apply (rule gmeet_cong_r[simplified], fact+)
  | 
| 
 | 
  3045  | 
  done
  | 
| 
 | 
  3046  | 
qed
  | 
| 
 | 
  3047  | 
  | 
| 
 | 
  3048  | 
(*
  | 
| 
 | 
  3049  | 
lemma (in gcd_condition_monoid) asc_cong_gcd_l [intro]:
  | 
| 
 | 
  3050  | 
  assumes carr: "b \<in> carrier G"
  | 
| 
 | 
  3051  | 
  shows "asc_cong (\<lambda>a. somegcd G a b)"
  | 
| 
 | 
  3052  | 
using carr
  | 
| 
 | 
  3053  | 
unfolding CONG_def
  | 
| 
 | 
  3054  | 
by clarsimp (blast intro: gcd_cong_l)
  | 
| 
 | 
  3055  | 
  | 
| 
 | 
  3056  | 
lemma (in gcd_condition_monoid) asc_cong_gcd_r [intro]:
  | 
| 
 | 
  3057  | 
  assumes carr: "a \<in> carrier G"
  | 
| 
 | 
  3058  | 
  shows "asc_cong (\<lambda>b. somegcd G a b)"
  | 
| 
 | 
  3059  | 
using carr
  | 
| 
 | 
  3060  | 
unfolding CONG_def
  | 
| 
 | 
  3061  | 
by clarsimp (blast intro: gcd_cong_r)
  | 
| 
 | 
  3062  | 
  | 
| 
 | 
  3063  | 
lemmas (in gcd_condition_monoid) asc_cong_gcd_split [simp] = 
  | 
| 
 | 
  3064  | 
    assoc_split[OF _ asc_cong_gcd_l] assoc_split[OF _ asc_cong_gcd_r]
  | 
| 
 | 
  3065  | 
*)
  | 
| 
 | 
  3066  | 
  | 
| 
 | 
  3067  | 
lemma (in gcd_condition_monoid) gcdI:
  | 
| 
 | 
  3068  | 
  assumes dvd: "a divides b"  "a divides c"
  | 
| 
 | 
  3069  | 
    and others: "\<forall>y\<in>carrier G. y divides b \<and> y divides c \<longrightarrow> y divides a"
  | 
| 
 | 
  3070  | 
    and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G" and ccarr: "c \<in> carrier G"
  | 
| 
 | 
  3071  | 
  shows "a \<sim> somegcd G b c"
  | 
| 
 | 
  3072  | 
apply (simp add: somegcd_def)
  | 
| 
 | 
  3073  | 
apply (rule someI2_ex)
  | 
| 
 | 
  3074  | 
 apply (rule exI[of _ a], simp add: isgcd_def)
  | 
| 
 | 
  3075  | 
 apply (simp add: assms)
  | 
| 
 | 
  3076  | 
apply (simp add: isgcd_def assms, clarify)
  | 
| 
 | 
  3077  | 
apply (insert assms, blast intro: associatedI)
  | 
| 
 | 
  3078  | 
done
  | 
| 
 | 
  3079  | 
  | 
| 
 | 
  3080  | 
lemma (in gcd_condition_monoid) gcdI2:
  | 
| 
 | 
  3081  | 
  assumes "a gcdof b c"
  | 
| 
 | 
  3082  | 
    and "a \<in> carrier G" and bcarr: "b \<in> carrier G" and ccarr: "c \<in> carrier G"
  | 
| 
 | 
  3083  | 
  shows "a \<sim> somegcd G b c"
  | 
| 
 | 
  3084  | 
using assms
  | 
| 
 | 
  3085  | 
unfolding isgcd_def
  | 
| 
 | 
  3086  | 
by (blast intro: gcdI)
  | 
| 
 | 
  3087  | 
  | 
| 
 | 
  3088  | 
lemma (in gcd_condition_monoid) SomeGcd_ex:
  | 
| 
 | 
  3089  | 
  assumes "finite A"  "A \<subseteq> carrier G"  "A \<noteq> {}"
 | 
| 
 | 
  3090  | 
  shows "\<exists>x\<in> carrier G. x = SomeGcd G A"
  | 
| 
 | 
  3091  | 
proof -
  | 
| 
 | 
  3092  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3093  | 
  show ?thesis
  | 
| 
 | 
  3094  | 
    apply (simp add: SomeGcd_def)
  | 
| 
 | 
  3095  | 
    apply (rule finite_ginf_closed[simplified], fact+)
  | 
| 
 | 
  3096  | 
  done
  | 
| 
 | 
  3097  | 
qed
  | 
| 
 | 
  3098  | 
  | 
| 
 | 
  3099  | 
lemma (in gcd_condition_monoid) gcd_assoc:
  | 
| 
 | 
  3100  | 
  assumes carr: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
  3101  | 
  shows "somegcd G (somegcd G a b) c \<sim> somegcd G a (somegcd G b c)"
  | 
| 
 | 
  3102  | 
proof -
  | 
| 
 | 
  3103  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3104  | 
  show ?thesis
  | 
| 
 | 
  3105  | 
    apply (subst (2 3) somegcd_gmeet, (simp add: carr)+)
  | 
| 
 | 
  3106  | 
    apply (simp add: somegcd_gmeet carr)
  | 
| 
 | 
  3107  | 
    apply (rule gmeet_assoc[simplified], fact+)
  | 
| 
 | 
  3108  | 
  done
  | 
| 
 | 
  3109  | 
qed
  | 
| 
 | 
  3110  | 
  | 
| 
 | 
  3111  | 
lemma (in gcd_condition_monoid) gcd_mult:
  | 
| 
 | 
  3112  | 
  assumes acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G" and ccarr: "c \<in> carrier G"
  | 
| 
 | 
  3113  | 
  shows "c \<otimes> somegcd G a b \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)"
  | 
| 
 | 
  3114  | 
proof - (* following Jacobson, Basic Algebra, p.140 *)
  | 
| 
 | 
  3115  | 
  let ?d = "somegcd G a b"
  | 
| 
 | 
  3116  | 
  let ?e = "somegcd G (c \<otimes> a) (c \<otimes> b)"
  | 
| 
 | 
  3117  | 
  note carr[simp] = acarr bcarr ccarr
  | 
| 
 | 
  3118  | 
  have dcarr: "?d \<in> carrier G" by simp
  | 
| 
 | 
  3119  | 
  have ecarr: "?e \<in> carrier G" by simp
  | 
| 
 | 
  3120  | 
  note carr = carr dcarr ecarr
  | 
| 
 | 
  3121  | 
  | 
| 
 | 
  3122  | 
  have "?d divides a" by (simp add: gcd_divides_l)
  | 
| 
 | 
  3123  | 
  hence cd'ca: "c \<otimes> ?d divides (c \<otimes> a)" by (simp add: divides_mult_lI)
  | 
| 
 | 
  3124  | 
  | 
| 
 | 
  3125  | 
  have "?d divides b" by (simp add: gcd_divides_r)
  | 
| 
 | 
  3126  | 
  hence cd'cb: "c \<otimes> ?d divides (c \<otimes> b)" by (simp add: divides_mult_lI)
  | 
| 
 | 
  3127  | 
  
  | 
| 
 | 
  3128  | 
  from cd'ca cd'cb
  | 
| 
 | 
  3129  | 
      have cd'e: "c \<otimes> ?d divides ?e"
  | 
| 
 | 
  3130  | 
      by (rule gcd_divides) simp+
  | 
| 
 | 
  3131  | 
  | 
| 
 | 
  3132  | 
  hence "\<exists>u. u \<in> carrier G \<and> ?e = c \<otimes> ?d \<otimes> u"
  | 
| 
 | 
  3133  | 
      by (elim dividesE, fast)
  | 
| 
 | 
  3134  | 
  from this obtain u
  | 
| 
 | 
  3135  | 
      where ucarr[simp]: "u \<in> carrier G"
  | 
| 
 | 
  3136  | 
      and e_cdu: "?e = c \<otimes> ?d \<otimes> u"
  | 
| 
 | 
  3137  | 
      by auto
  | 
| 
 | 
  3138  | 
  | 
| 
 | 
  3139  | 
  note carr = carr ucarr
  | 
| 
 | 
  3140  | 
  | 
| 
 | 
  3141  | 
  have "?e divides c \<otimes> a" by (rule gcd_divides_l) simp+
  | 
| 
 | 
  3142  | 
  hence "\<exists>x. x \<in> carrier G \<and> c \<otimes> a = ?e \<otimes> x"
  | 
| 
 | 
  3143  | 
      by (elim dividesE, fast)
  | 
| 
 | 
  3144  | 
  from this obtain x
  | 
| 
 | 
  3145  | 
      where xcarr: "x \<in> carrier G"
  | 
| 
 | 
  3146  | 
      and ca_ex: "c \<otimes> a = ?e \<otimes> x"
  | 
| 
 | 
  3147  | 
      by auto
  | 
| 
 | 
  3148  | 
  with e_cdu
  | 
| 
 | 
  3149  | 
      have ca_cdux: "c \<otimes> a = c \<otimes> ?d \<otimes> u \<otimes> x" by simp
  | 
| 
 | 
  3150  | 
  | 
| 
 | 
  3151  | 
  from ca_cdux xcarr
  | 
| 
 | 
  3152  | 
       have "c \<otimes> a = c \<otimes> (?d \<otimes> u \<otimes> x)" by (simp add: m_assoc)
  | 
| 
 | 
  3153  | 
  then have "a = ?d \<otimes> u \<otimes> x" by (rule l_cancel[of c a]) (simp add: xcarr)+
  | 
| 
 | 
  3154  | 
  hence du'a: "?d \<otimes> u divides a" by (rule dividesI[OF xcarr])
  | 
| 
 | 
  3155  | 
  | 
| 
 | 
  3156  | 
  have "?e divides c \<otimes> b" by (intro gcd_divides_r, simp+)
  | 
| 
 | 
  3157  | 
  hence "\<exists>x. x \<in> carrier G \<and> c \<otimes> b = ?e \<otimes> x"
  | 
| 
 | 
  3158  | 
      by (elim dividesE, fast)
  | 
| 
 | 
  3159  | 
  from this obtain x
  | 
| 
 | 
  3160  | 
      where xcarr: "x \<in> carrier G"
  | 
| 
 | 
  3161  | 
      and cb_ex: "c \<otimes> b = ?e \<otimes> x"
  | 
| 
 | 
  3162  | 
      by auto
  | 
| 
 | 
  3163  | 
  with e_cdu
  | 
| 
 | 
  3164  | 
      have cb_cdux: "c \<otimes> b = c \<otimes> ?d \<otimes> u \<otimes> x" by simp
  | 
| 
 | 
  3165  | 
  | 
| 
 | 
  3166  | 
  from cb_cdux xcarr
  | 
| 
 | 
  3167  | 
      have "c \<otimes> b = c \<otimes> (?d \<otimes> u \<otimes> x)" by (simp add: m_assoc)
  | 
| 
 | 
  3168  | 
  with xcarr
  | 
| 
 | 
  3169  | 
      have "b = ?d \<otimes> u \<otimes> x" by (intro l_cancel[of c b], simp+)
  | 
| 
 | 
  3170  | 
  hence du'b: "?d \<otimes> u divides b" by (intro dividesI[OF xcarr])
  | 
| 
 | 
  3171  | 
  | 
| 
 | 
  3172  | 
  from du'a du'b carr
  | 
| 
 | 
  3173  | 
      have du'd: "?d \<otimes> u divides ?d"
  | 
| 
 | 
  3174  | 
      by (intro gcd_divides, simp+)
  | 
| 
 | 
  3175  | 
  hence uunit: "u \<in> Units G"
  | 
| 
 | 
  3176  | 
  proof (elim dividesE)
  | 
| 
 | 
  3177  | 
    fix v
  | 
| 
 | 
  3178  | 
    assume vcarr[simp]: "v \<in> carrier G"
  | 
| 
 | 
  3179  | 
    assume d: "?d = ?d \<otimes> u \<otimes> v"
  | 
| 
 | 
  3180  | 
    have "?d \<otimes> \<one> = ?d \<otimes> u \<otimes> v" by simp fact
  | 
| 
 | 
  3181  | 
    also have "?d \<otimes> u \<otimes> v = ?d \<otimes> (u \<otimes> v)" by (simp add: m_assoc)
  | 
| 
 | 
  3182  | 
    finally have "?d \<otimes> \<one> = ?d \<otimes> (u \<otimes> v)" .
  | 
| 
 | 
  3183  | 
    hence i2: "\<one> = u \<otimes> v" by (rule l_cancel) simp+
  | 
| 
 | 
  3184  | 
    hence i1: "\<one> = v \<otimes> u" by (simp add: m_comm)
  | 
| 
 | 
  3185  | 
    from vcarr i1[symmetric] i2[symmetric]
  | 
| 
 | 
  3186  | 
        show "u \<in> Units G"
  | 
| 
 | 
  3187  | 
        by (unfold Units_def, simp, fast)
  | 
| 
 | 
  3188  | 
  qed
  | 
| 
 | 
  3189  | 
  | 
| 
 | 
  3190  | 
  from e_cdu uunit
  | 
| 
 | 
  3191  | 
      have "somegcd G (c \<otimes> a) (c \<otimes> b) \<sim> c \<otimes> somegcd G a b"
  | 
| 
 | 
  3192  | 
      by (intro associatedI2[of u], simp+)
  | 
| 
 | 
  3193  | 
  from this[symmetric]
  | 
| 
 | 
  3194  | 
      show "c \<otimes> somegcd G a b \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
  | 
| 
 | 
  3195  | 
qed
  | 
| 
 | 
  3196  | 
  | 
| 
 | 
  3197  | 
lemma (in monoid) assoc_subst:
  | 
| 
 | 
  3198  | 
  assumes ab: "a \<sim> b"
  | 
| 
 | 
  3199  | 
    and cP: "ALL a b. a : carrier G & b : carrier G & a \<sim> b
  | 
| 
 | 
  3200  | 
      --> f a : carrier G & f b : carrier G & f a \<sim> f b"
  | 
| 
 | 
  3201  | 
    and carr: "a \<in> carrier G"  "b \<in> carrier G"
  | 
| 
 | 
  3202  | 
  shows "f a \<sim> f b"
  | 
| 
 | 
  3203  | 
  using assms by auto
  | 
| 
 | 
  3204  | 
  | 
| 
 | 
  3205  | 
lemma (in gcd_condition_monoid) relprime_mult:
  | 
| 
 | 
  3206  | 
  assumes abrelprime: "somegcd G a b \<sim> \<one>" and acrelprime: "somegcd G a c \<sim> \<one>"
  | 
| 
 | 
  3207  | 
    and carr[simp]: "a \<in> carrier G"  "b \<in> carrier G"  "c \<in> carrier G"
  | 
| 
 | 
  3208  | 
  shows "somegcd G a (b \<otimes> c) \<sim> \<one>"
  | 
| 
 | 
  3209  | 
proof -
  | 
| 
 | 
  3210  | 
  have "c = c \<otimes> \<one>" by simp
  | 
| 
 | 
  3211  | 
  also from abrelprime[symmetric]
  | 
| 
 | 
  3212  | 
       have "\<dots> \<sim> c \<otimes> somegcd G a b"
  | 
| 
 | 
  3213  | 
	 by (rule assoc_subst) (simp add: mult_cong_r)+
  | 
| 
 | 
  3214  | 
  also have "\<dots> \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by (rule gcd_mult) fact+
  | 
| 
 | 
  3215  | 
  finally
  | 
| 
 | 
  3216  | 
       have c: "c \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
  | 
| 
 | 
  3217  | 
  | 
| 
 | 
  3218  | 
  from carr
  | 
| 
 | 
  3219  | 
       have a: "a \<sim> somegcd G a (c \<otimes> a)"
  | 
| 
 | 
  3220  | 
       by (fast intro: gcdI divides_prod_l)
  | 
| 
 | 
  3221  | 
  | 
| 
 | 
  3222  | 
  have "somegcd G a (b \<otimes> c) \<sim> somegcd G a (c \<otimes> b)" by (simp add: m_comm)
  | 
| 
 | 
  3223  | 
  also from a
  | 
| 
 | 
  3224  | 
       have "\<dots> \<sim> somegcd G (somegcd G a (c \<otimes> a)) (c \<otimes> b)"
  | 
| 
 | 
  3225  | 
	 by (rule assoc_subst) (simp add: gcd_cong_l)+
  | 
| 
 | 
  3226  | 
  also from gcd_assoc
  | 
| 
 | 
  3227  | 
       have "\<dots> \<sim> somegcd G a (somegcd G (c \<otimes> a) (c \<otimes> b))"
  | 
| 
 | 
  3228  | 
       by (rule assoc_subst) simp+
  | 
| 
 | 
  3229  | 
  also from c[symmetric]
  | 
| 
 | 
  3230  | 
       have "\<dots> \<sim> somegcd G a c"
  | 
| 
 | 
  3231  | 
	 by (rule assoc_subst) (simp add: gcd_cong_r)+
  | 
| 
 | 
  3232  | 
  also note acrelprime
  | 
| 
 | 
  3233  | 
  finally
  | 
| 
 | 
  3234  | 
       show "somegcd G a (b \<otimes> c) \<sim> \<one>" by simp
  | 
| 
 | 
  3235  | 
qed
  | 
| 
 | 
  3236  | 
  | 
| 
 | 
  3237  | 
lemma (in gcd_condition_monoid) primeness_condition:
  | 
| 
 | 
  3238  | 
  "primeness_condition_monoid G"
  | 
| 
 | 
  3239  | 
apply unfold_locales
  | 
| 
 | 
  3240  | 
apply (rule primeI)
  | 
| 
 | 
  3241  | 
 apply (elim irreducibleE, assumption)
  | 
| 
 | 
  3242  | 
proof -
  | 
| 
 | 
  3243  | 
  fix p a b
  | 
| 
 | 
  3244  | 
  assume pcarr: "p \<in> carrier G" and acarr: "a \<in> carrier G" and bcarr: "b \<in> carrier G"
  | 
| 
 | 
  3245  | 
    and pirr: "irreducible G p"
  | 
| 
 | 
  3246  | 
    and pdvdab: "p divides a \<otimes> b"
  | 
| 
 | 
  3247  | 
  from pirr
  | 
| 
 | 
  3248  | 
      have pnunit: "p \<notin> Units G"
  | 
| 
 | 
  3249  | 
      and r[rule_format]: "\<forall>b. b \<in> carrier G \<and> properfactor G b p \<longrightarrow> b \<in> Units G"
  | 
| 
 | 
  3250  | 
      by - (fast elim: irreducibleE)+
  | 
| 
 | 
  3251  | 
  | 
| 
 | 
  3252  | 
  show "p divides a \<or> p divides b"
  | 
| 
 | 
  3253  | 
  proof (rule ccontr, clarsimp)
  | 
| 
 | 
  3254  | 
    assume npdvda: "\<not> p divides a"
  | 
| 
 | 
  3255  | 
    with pcarr acarr
  | 
| 
 | 
  3256  | 
    have "\<one> \<sim> somegcd G p a"
  | 
| 
 | 
  3257  | 
    apply (intro gcdI, simp, simp, simp)
  | 
| 
 | 
  3258  | 
      apply (fast intro: unit_divides)
  | 
| 
 | 
  3259  | 
     apply (fast intro: unit_divides)
  | 
| 
 | 
  3260  | 
    apply (clarsimp simp add: Unit_eq_dividesone[symmetric])
  | 
| 
 | 
  3261  | 
    apply (rule r, rule, assumption)
  | 
| 
 | 
  3262  | 
    apply (rule properfactorI, assumption)
  | 
| 
 | 
  3263  | 
    proof (rule ccontr, simp)
  | 
| 
 | 
  3264  | 
      fix y
  | 
| 
 | 
  3265  | 
      assume ycarr: "y \<in> carrier G"
  | 
| 
 | 
  3266  | 
      assume "p divides y"
  | 
| 
 | 
  3267  | 
      also assume "y divides a"
  | 
| 
 | 
  3268  | 
      finally
  | 
| 
 | 
  3269  | 
          have "p divides a" by (simp add: pcarr ycarr acarr)
  | 
| 
 | 
  3270  | 
      with npdvda
  | 
| 
 | 
  3271  | 
          show "False" ..
  | 
| 
 | 
  3272  | 
    qed simp+
  | 
| 
 | 
  3273  | 
    with pcarr acarr
  | 
| 
 | 
  3274  | 
        have pa: "somegcd G p a \<sim> \<one>" by (fast intro: associated_sym[of "\<one>"] gcd_closed)
  | 
| 
 | 
  3275  | 
  | 
| 
 | 
  3276  | 
    assume npdvdb: "\<not> p divides b"
  | 
| 
 | 
  3277  | 
    with pcarr bcarr
  | 
| 
 | 
  3278  | 
    have "\<one> \<sim> somegcd G p b"
  | 
| 
 | 
  3279  | 
    apply (intro gcdI, simp, simp, simp)
  | 
| 
 | 
  3280  | 
      apply (fast intro: unit_divides)
  | 
| 
 | 
  3281  | 
     apply (fast intro: unit_divides)
  | 
| 
 | 
  3282  | 
    apply (clarsimp simp add: Unit_eq_dividesone[symmetric])
  | 
| 
 | 
  3283  | 
    apply (rule r, rule, assumption)
  | 
| 
 | 
  3284  | 
    apply (rule properfactorI, assumption)
  | 
| 
 | 
  3285  | 
    proof (rule ccontr, simp)
  | 
| 
 | 
  3286  | 
      fix y
  | 
| 
 | 
  3287  | 
      assume ycarr: "y \<in> carrier G"
  | 
| 
 | 
  3288  | 
      assume "p divides y"
  | 
| 
 | 
  3289  | 
      also assume "y divides b"
  | 
| 
 | 
  3290  | 
      finally have "p divides b" by (simp add: pcarr ycarr bcarr)
  | 
| 
 | 
  3291  | 
      with npdvdb
  | 
| 
 | 
  3292  | 
          show "False" ..
  | 
| 
 | 
  3293  | 
    qed simp+
  | 
| 
 | 
  3294  | 
    with pcarr bcarr
  | 
| 
 | 
  3295  | 
        have pb: "somegcd G p b \<sim> \<one>" by (fast intro: associated_sym[of "\<one>"] gcd_closed)
  | 
| 
 | 
  3296  | 
  | 
| 
 | 
  3297  | 
    from pcarr acarr bcarr pdvdab
  | 
| 
 | 
  3298  | 
        have "p gcdof p (a \<otimes> b)" by (fast intro: isgcd_divides_l)
  | 
| 
 | 
  3299  | 
  | 
| 
 | 
  3300  | 
    with pcarr acarr bcarr
  | 
| 
 | 
  3301  | 
         have "p \<sim> somegcd G p (a \<otimes> b)" by (fast intro: gcdI2)
  | 
| 
 | 
  3302  | 
    also from pa pb pcarr acarr bcarr
  | 
| 
 | 
  3303  | 
         have "somegcd G p (a \<otimes> b) \<sim> \<one>" by (rule relprime_mult)
  | 
| 
 | 
  3304  | 
    finally have "p \<sim> \<one>" by (simp add: pcarr acarr bcarr)
  | 
| 
 | 
  3305  | 
  | 
| 
 | 
  3306  | 
    with pcarr
  | 
| 
 | 
  3307  | 
        have "p \<in> Units G" by (fast intro: assoc_unit_l)
  | 
| 
 | 
  3308  | 
    with pnunit
  | 
| 
 | 
  3309  | 
        show "False" ..
  | 
| 
 | 
  3310  | 
  qed
  | 
| 
 | 
  3311  | 
qed
  | 
| 
 | 
  3312  | 
  | 
| 
 | 
  3313  | 
interpretation gcd_condition_monoid \<subseteq> primeness_condition_monoid
  | 
| 
 | 
  3314  | 
  by (rule primeness_condition)
  | 
| 
 | 
  3315  | 
  | 
| 
 | 
  3316  | 
  | 
| 
 | 
  3317  | 
subsubsection {* Divisor chain condition *}
 | 
| 
 | 
  3318  | 
  | 
| 
 | 
  3319  | 
lemma (in divisor_chain_condition_monoid) wfactors_exist:
  | 
| 
 | 
  3320  | 
  assumes acarr: "a \<in> carrier G"
  | 
| 
 | 
  3321  | 
  shows "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a"
  | 
| 
 | 
  3322  | 
proof -
  | 
| 
 | 
  3323  | 
  have r[rule_format]: "a \<in> carrier G \<longrightarrow> (\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a)"
  | 
| 
 | 
  3324  | 
    apply (rule wf_induct[OF division_wellfounded])
  | 
| 
 | 
  3325  | 
  proof -
  | 
| 
 | 
  3326  | 
    fix x
  | 
| 
 | 
  3327  | 
    assume ih: "\<forall>y. (y, x) \<in> {(x, y). x \<in> carrier G \<and> y \<in> carrier G \<and> properfactor G x y}
 | 
| 
 | 
  3328  | 
                    \<longrightarrow> y \<in> carrier G \<longrightarrow> (\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as y)"
  | 
| 
 | 
  3329  | 
  | 
| 
 | 
  3330  | 
    show "x \<in> carrier G \<longrightarrow> (\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as x)"
  | 
| 
 | 
  3331  | 
    apply clarify
  | 
| 
 | 
  3332  | 
    apply (cases "x \<in> Units G")
  | 
| 
 | 
  3333  | 
     apply (rule exI[of _ "[]"], simp)
  | 
| 
 | 
  3334  | 
    apply (cases "irreducible G x")
  | 
| 
 | 
  3335  | 
     apply (rule exI[of _ "[x]"], simp add: wfactors_def)
  | 
| 
 | 
  3336  | 
    proof -
  | 
| 
 | 
  3337  | 
      assume xcarr: "x \<in> carrier G"
  | 
| 
 | 
  3338  | 
        and xnunit: "x \<notin> Units G"
  | 
| 
 | 
  3339  | 
        and xnirr: "\<not> irreducible G x"
  | 
| 
 | 
  3340  | 
      hence "\<exists>y. y \<in> carrier G \<and> properfactor G y x \<and> y \<notin> Units G"
  | 
| 
 | 
  3341  | 
        apply - apply (rule ccontr, simp)
  | 
| 
 | 
  3342  | 
        apply (subgoal_tac "irreducible G x", simp)
  | 
| 
 | 
  3343  | 
        apply (rule irreducibleI, simp, simp)
  | 
| 
 | 
  3344  | 
      done
  | 
| 
 | 
  3345  | 
      from this obtain y
  | 
| 
 | 
  3346  | 
          where ycarr: "y \<in> carrier G"
  | 
| 
 | 
  3347  | 
          and ynunit: "y \<notin> Units G"
  | 
| 
 | 
  3348  | 
          and pfyx: "properfactor G y x"
  | 
| 
 | 
  3349  | 
          by auto
  | 
| 
 | 
  3350  | 
  | 
| 
 | 
  3351  | 
      have ih':
  | 
| 
 | 
  3352  | 
           "\<And>y. \<lbrakk>y \<in> carrier G; properfactor G y x\<rbrakk>
  | 
| 
 | 
  3353  | 
                \<Longrightarrow> \<exists>as. set as \<subseteq> carrier G \<and> wfactors G as y"
  | 
| 
 | 
  3354  | 
          by (rule ih[rule_format, simplified]) (simp add: xcarr)+
  | 
| 
 | 
  3355  | 
  | 
| 
 | 
  3356  | 
      from ycarr pfyx
  | 
| 
 | 
  3357  | 
          have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as y"
  | 
| 
 | 
  3358  | 
          by (rule ih')
  | 
| 
 | 
  3359  | 
      from this obtain ys
  | 
| 
 | 
  3360  | 
          where yscarr: "set ys \<subseteq> carrier G"
  | 
| 
 | 
  3361  | 
          and yfs: "wfactors G ys y"
  | 
| 
 | 
  3362  | 
          by auto
  | 
| 
 | 
  3363  | 
  | 
| 
 | 
  3364  | 
      from pfyx
  | 
| 
 | 
  3365  | 
          have "y divides x"
  | 
| 
 | 
  3366  | 
          and nyx: "\<not> y \<sim> x"
  | 
| 
 | 
  3367  | 
          by - (fast elim: properfactorE2)+
  | 
| 
 | 
  3368  | 
      hence "\<exists>z. z \<in> carrier G \<and> x = y \<otimes> z"
  | 
| 
 | 
  3369  | 
          by (fast elim: dividesE)
  | 
| 
 | 
  3370  | 
  | 
| 
 | 
  3371  | 
      from this obtain z
  | 
| 
 | 
  3372  | 
          where zcarr: "z \<in> carrier G"
  | 
| 
 | 
  3373  | 
          and x: "x = y \<otimes> z"
  | 
| 
 | 
  3374  | 
          by auto
  | 
| 
 | 
  3375  | 
  | 
| 
 | 
  3376  | 
      from zcarr ycarr
  | 
| 
 | 
  3377  | 
      have "properfactor G z x"
  | 
| 
 | 
  3378  | 
        apply (subst x)
  | 
| 
 | 
  3379  | 
        apply (intro properfactorI3[of _ _ y])
  | 
| 
 | 
  3380  | 
         apply (simp add: m_comm)
  | 
| 
 | 
  3381  | 
        apply (simp add: ynunit)+
  | 
| 
 | 
  3382  | 
      done
  | 
| 
 | 
  3383  | 
      with zcarr
  | 
| 
 | 
  3384  | 
          have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as z"
  | 
| 
 | 
  3385  | 
          by (rule ih')
  | 
| 
 | 
  3386  | 
      from this obtain zs
  | 
| 
 | 
  3387  | 
          where zscarr: "set zs \<subseteq> carrier G"
  | 
| 
 | 
  3388  | 
          and zfs: "wfactors G zs z"
  | 
| 
 | 
  3389  | 
          by auto
  | 
| 
 | 
  3390  | 
  | 
| 
 | 
  3391  | 
      from yscarr zscarr
  | 
| 
 | 
  3392  | 
          have xscarr: "set (ys@zs) \<subseteq> carrier G" by simp
  | 
| 
 | 
  3393  | 
      from yfs zfs ycarr zcarr yscarr zscarr
  | 
| 
 | 
  3394  | 
          have "wfactors G (ys@zs) (y\<otimes>z)" by (rule wfactors_mult)
  | 
| 
 | 
  3395  | 
      hence "wfactors G (ys@zs) x" by (simp add: x)
  | 
| 
 | 
  3396  | 
  | 
| 
 | 
  3397  | 
      from xscarr this
  | 
| 
 | 
  3398  | 
          show "\<exists>xs. set xs \<subseteq> carrier G \<and> wfactors G xs x" by fast
  | 
| 
 | 
  3399  | 
    qed
  | 
| 
 | 
  3400  | 
  qed
  | 
| 
 | 
  3401  | 
  | 
| 
 | 
  3402  | 
  from acarr
  | 
| 
 | 
  3403  | 
      show ?thesis by (rule r)
  | 
| 
 | 
  3404  | 
qed
  | 
| 
 | 
  3405  | 
  | 
| 
 | 
  3406  | 
  | 
| 
 | 
  3407  | 
subsubsection {* Primeness condition *}
 | 
| 
 | 
  3408  | 
  | 
| 
 | 
  3409  | 
lemma (in comm_monoid_cancel) multlist_prime_pos:
  | 
| 
 | 
  3410  | 
  assumes carr: "a \<in> carrier G"  "set as \<subseteq> carrier G"
  | 
| 
 | 
  3411  | 
    and aprime: "prime G a"
  | 
| 
 | 
  3412  | 
    and "a divides (foldr (op \<otimes>) as \<one>)"
  | 
| 
 | 
  3413  | 
  shows "\<exists>i<length as. a divides (as!i)"
  | 
| 
 | 
  3414  | 
proof -
  | 
| 
 | 
  3415  | 
  have r[rule_format]:
  | 
| 
 | 
  3416  | 
       "set as \<subseteq> carrier G \<and> a divides (foldr (op \<otimes>) as \<one>)
  | 
| 
 | 
  3417  | 
        \<longrightarrow> (\<exists>i. i < length as \<and> a divides (as!i))"
  | 
| 
 | 
  3418  | 
    apply (induct as)
  | 
| 
 | 
  3419  | 
     apply clarsimp defer 1
  | 
| 
 | 
  3420  | 
     apply clarsimp defer 1
  | 
| 
 | 
  3421  | 
  proof -
  | 
| 
 | 
  3422  | 
    assume "a divides \<one>"
  | 
| 
 | 
  3423  | 
    with carr
  | 
| 
 | 
  3424  | 
        have "a \<in> Units G"
  | 
| 
 | 
  3425  | 
        by (fast intro: divides_unit[of a \<one>])
  | 
| 
 | 
  3426  | 
    with aprime
  | 
| 
 | 
  3427  | 
        show "False" by (elim primeE, simp)
  | 
| 
 | 
  3428  | 
  next
  | 
| 
 | 
  3429  | 
    fix aa as
  | 
| 
 | 
  3430  | 
    assume ih[rule_format]: "a divides foldr op \<otimes> as \<one> \<longrightarrow> (\<exists>i<length as. a divides as ! i)"
  | 
| 
 | 
  3431  | 
      and carr': "aa \<in> carrier G"  "set as \<subseteq> carrier G"
  | 
| 
 | 
  3432  | 
      and "a divides aa \<otimes> foldr op \<otimes> as \<one>"
  | 
| 
 | 
  3433  | 
    with carr aprime
  | 
| 
 | 
  3434  | 
        have "a divides aa \<or> a divides foldr op \<otimes> as \<one>"
  | 
| 
 | 
  3435  | 
        by (intro prime_divides) simp+
  | 
| 
 | 
  3436  | 
    moreover {
 | 
| 
 | 
  3437  | 
      assume "a divides aa"
  | 
| 
 | 
  3438  | 
      hence p1: "a divides (aa#as)!0" by simp
  | 
| 
 | 
  3439  | 
      have "0 < Suc (length as)" by simp
  | 
| 
 | 
  3440  | 
      with p1 have "\<exists>i<Suc (length as). a divides (aa # as) ! i" by fast
  | 
| 
 | 
  3441  | 
    }
  | 
| 
 | 
  3442  | 
    moreover {
 | 
| 
 | 
  3443  | 
      assume "a divides foldr op \<otimes> as \<one>"
  | 
| 
 | 
  3444  | 
      hence "\<exists>i. i < length as \<and> a divides as ! i" by (rule ih)
  | 
| 
 | 
  3445  | 
      from this obtain i where "a divides as ! i" and len: "i < length as" by auto
  | 
| 
 | 
  3446  | 
      hence p1: "a divides (aa#as) ! (Suc i)" by simp
  | 
| 
 | 
  3447  | 
      from len have "Suc i < Suc (length as)" by simp
  | 
| 
 | 
  3448  | 
      with p1 have "\<exists>i<Suc (length as). a divides (aa # as) ! i" by force
  | 
| 
 | 
  3449  | 
   }
  | 
| 
 | 
  3450  | 
   ultimately
  | 
| 
 | 
  3451  | 
      show "\<exists>i<Suc (length as). a divides (aa # as) ! i" by fast
  | 
| 
 | 
  3452  | 
  qed
  | 
| 
 | 
  3453  | 
  | 
| 
 | 
  3454  | 
  from assms
  | 
| 
 | 
  3455  | 
      show ?thesis
  | 
| 
 | 
  3456  | 
      by (intro r, safe)
  | 
| 
 | 
  3457  | 
qed
  | 
| 
 | 
  3458  | 
  | 
| 
 | 
  3459  | 
lemma (in primeness_condition_monoid) wfactors_unique__hlp_induct:
  | 
| 
 | 
  3460  | 
  "\<forall>a as'. a \<in> carrier G \<and> set as \<subseteq> carrier G \<and> set as' \<subseteq> carrier G \<and> 
  | 
| 
 | 
  3461  | 
           wfactors G as a \<and> wfactors G as' a \<longrightarrow> essentially_equal G as as'"
  | 
| 
 | 
  3462  | 
apply (induct as)
  | 
| 
 | 
  3463  | 
apply clarsimp defer 1
  | 
| 
 | 
  3464  | 
apply clarsimp defer 1
  | 
| 
 | 
  3465  | 
proof -
  | 
| 
 | 
  3466  | 
  fix a as'
  | 
| 
 | 
  3467  | 
  assume acarr: "a \<in> carrier G"
  | 
| 
 | 
  3468  | 
    and "wfactors G [] a"
  | 
| 
 | 
  3469  | 
  hence aunit: "a \<in> Units G"
  | 
| 
 | 
  3470  | 
    apply (elim wfactorsE)
  | 
| 
 | 
  3471  | 
    apply (simp, rule assoc_unit_r[of "\<one>"], simp+)
  | 
| 
 | 
  3472  | 
  done
  | 
| 
 | 
  3473  | 
  | 
| 
 | 
  3474  | 
  assume "set as' \<subseteq> carrier G"  "wfactors G as' a"
  | 
| 
 | 
  3475  | 
  with aunit
  | 
| 
 | 
  3476  | 
      have "as' = []"
  | 
| 
 | 
  3477  | 
      by (intro unit_wfactors_empty[of a])
  | 
| 
 | 
  3478  | 
  thus "essentially_equal G [] as'" by simp
  | 
| 
 | 
  3479  | 
next
  | 
| 
 | 
  3480  | 
  fix a as ah as'
  | 
| 
 | 
  3481  | 
  assume ih[rule_format]: 
  | 
| 
 | 
  3482  | 
         "\<forall>a as'. a \<in> carrier G \<and> set as' \<subseteq> carrier G \<and> 
  | 
| 
 | 
  3483  | 
                  wfactors G as a \<and> wfactors G as' a \<longrightarrow> essentially_equal G as as'"
  | 
| 
 | 
  3484  | 
    and acarr: "a \<in> carrier G" and ahcarr: "ah \<in> carrier G"
  | 
| 
 | 
  3485  | 
    and ascarr: "set as \<subseteq> carrier G" and as'carr: "set as' \<subseteq> carrier G"
  | 
| 
 | 
  3486  | 
    and afs: "wfactors G (ah # as) a"
  | 
| 
 | 
  3487  | 
    and afs': "wfactors G as' a"
  | 
| 
 | 
  3488  | 
  hence ahdvda: "ah divides a"
  | 
| 
 | 
  3489  | 
      by (intro wfactors_dividesI[of "ah#as" "a"], simp+)
  | 
| 
 | 
  3490  | 
  hence "\<exists>a'\<in> carrier G. a = ah \<otimes> a'" by (fast elim: dividesE)
  | 
| 
 | 
  3491  | 
  from this obtain a'
  | 
| 
 | 
  3492  | 
      where a'carr: "a' \<in> carrier G"
  | 
| 
 | 
  3493  | 
      and a: "a = ah \<otimes> a'"
  | 
| 
 | 
  3494  | 
      by auto
  | 
| 
 | 
  3495  | 
  have a'fs: "wfactors G as a'"
  | 
| 
 | 
  3496  | 
    apply (rule wfactorsE[OF afs], rule wfactorsI, simp)
  | 
| 
 | 
  3497  | 
    apply (simp add: a, insert ascarr a'carr)
  | 
| 
 | 
  3498  | 
    apply (intro assoc_l_cancel[of ah _ a'] multlist_closed ahcarr, assumption+)
  | 
| 
 | 
  3499  | 
  done
  | 
| 
 | 
  3500  | 
  | 
| 
 | 
  3501  | 
  from afs have ahirr: "irreducible G ah" by (elim wfactorsE, simp)
  | 
| 
 | 
  3502  | 
  with ascarr have ahprime: "prime G ah" by (intro irreducible_prime ahcarr)
  | 
| 
 | 
  3503  | 
  | 
| 
 | 
  3504  | 
  note carr [simp] = acarr ahcarr ascarr as'carr a'carr
  | 
| 
 | 
  3505  | 
  | 
| 
 | 
  3506  | 
  note ahdvda
  | 
| 
 | 
  3507  | 
  also from afs'
  | 
| 
 | 
  3508  | 
       have "a divides (foldr (op \<otimes>) as' \<one>)"
  | 
| 
 | 
  3509  | 
       by (elim wfactorsE associatedE, simp)
  | 
| 
 | 
  3510  | 
  finally have "ah divides (foldr (op \<otimes>) as' \<one>)" by simp
  | 
| 
 | 
  3511  | 
  | 
| 
 | 
  3512  | 
  with ahprime
  | 
| 
 | 
  3513  | 
      have "\<exists>i<length as'. ah divides as'!i"
  | 
| 
 | 
  3514  | 
      by (intro multlist_prime_pos, simp+)
  | 
| 
 | 
  3515  | 
  from this obtain i
  | 
| 
 | 
  3516  | 
      where len: "i<length as'" and ahdvd: "ah divides as'!i"
  | 
| 
 | 
  3517  | 
      by auto
  | 
| 
 | 
  3518  | 
  from afs' carr have irrasi: "irreducible G (as'!i)"
  | 
| 
 | 
  3519  | 
      by (fast intro: nth_mem[OF len] elim: wfactorsE)
  | 
| 
 | 
  3520  | 
  from len carr
  | 
| 
 | 
  3521  | 
      have asicarr[simp]: "as'!i \<in> carrier G" by (unfold set_conv_nth, force)
  | 
| 
 | 
  3522  | 
  note carr = carr asicarr
  | 
| 
 | 
  3523  | 
  | 
| 
 | 
  3524  | 
  from ahdvd have "\<exists>x \<in> carrier G. as'!i = ah \<otimes> x" by (fast elim: dividesE)
  | 
| 
 | 
  3525  | 
  from this obtain x where "x \<in> carrier G" and asi: "as'!i = ah \<otimes> x" by auto
  | 
| 
 | 
  3526  | 
  | 
| 
 | 
  3527  | 
  with carr irrasi[simplified asi]
  | 
| 
 | 
  3528  | 
      have asiah: "as'!i \<sim> ah" apply -
  | 
| 
 | 
  3529  | 
    apply (elim irreducible_prodE[of "ah" "x"], assumption+)
  | 
| 
 | 
  3530  | 
     apply (rule associatedI2[of x], assumption+)
  | 
| 
 | 
  3531  | 
    apply (rule irreducibleE[OF ahirr], simp)
  | 
| 
 | 
  3532  | 
  done
  | 
| 
 | 
  3533  | 
  | 
| 
 | 
  3534  | 
  note setparts = set_take_subset[of i as'] set_drop_subset[of "Suc i" as']
  | 
| 
 | 
  3535  | 
  note partscarr [simp] = setparts[THEN subset_trans[OF _ as'carr]]
  | 
| 
 | 
  3536  | 
  note carr = carr partscarr
  | 
| 
 | 
  3537  | 
  | 
| 
 | 
  3538  | 
  have "\<exists>aa_1. aa_1 \<in> carrier G \<and> wfactors G (take i as') aa_1"
  | 
| 
 | 
  3539  | 
    apply (intro wfactors_prod_exists)
  | 
| 
 | 
  3540  | 
    using setparts afs' by (fast elim: wfactorsE, simp)
  | 
| 
 | 
  3541  | 
  from this obtain aa_1
  | 
| 
 | 
  3542  | 
      where aa1carr: "aa_1 \<in> carrier G"
  | 
| 
 | 
  3543  | 
      and aa1fs: "wfactors G (take i as') aa_1"
  | 
| 
 | 
  3544  | 
      by auto
  | 
| 
 | 
  3545  | 
  | 
| 
 | 
  3546  | 
  have "\<exists>aa_2. aa_2 \<in> carrier G \<and> wfactors G (drop (Suc i) as') aa_2"
  | 
| 
 | 
  3547  | 
    apply (intro wfactors_prod_exists)
  | 
| 
 | 
  3548  | 
    using setparts afs' by (fast elim: wfactorsE, simp)
  | 
| 
 | 
  3549  | 
  from this obtain aa_2
  | 
| 
 | 
  3550  | 
      where aa2carr: "aa_2 \<in> carrier G"
  | 
| 
 | 
  3551  | 
      and aa2fs: "wfactors G (drop (Suc i) as') aa_2"
  | 
| 
 | 
  3552  | 
      by auto
  | 
| 
 | 
  3553  | 
  | 
| 
 | 
  3554  | 
  note carr = carr aa1carr[simp] aa2carr[simp]
  | 
| 
 | 
  3555  | 
  | 
| 
 | 
  3556  | 
  from aa1fs aa2fs
  | 
| 
 | 
  3557  | 
      have v1: "wfactors G (take i as' @ drop (Suc i) as') (aa_1 \<otimes> aa_2)"
  | 
| 
 | 
  3558  | 
      by (intro wfactors_mult, simp+)
  | 
| 
 | 
  3559  | 
  hence v1': "wfactors G (as'!i # take i as' @ drop (Suc i) as') (as'!i \<otimes> (aa_1 \<otimes> aa_2))"
  | 
| 
 | 
  3560  | 
      apply (intro wfactors_mult_single)
  | 
| 
 | 
  3561  | 
      using setparts afs'
  | 
| 
 | 
  3562  | 
      by (fast intro: nth_mem[OF len] elim: wfactorsE, simp+)
  | 
| 
 | 
  3563  | 
  | 
| 
 | 
  3564  | 
  from aa2carr carr aa1fs aa2fs
  | 
| 
 | 
  3565  | 
      have "wfactors G (as'!i # drop (Suc i) as') (as'!i \<otimes> aa_2)"
  | 
| 
 | 
  3566  | 
    apply (intro wfactors_mult_single)
  | 
| 
 | 
  3567  | 
        apply (rule wfactorsE[OF afs'], fast intro: nth_mem[OF len])
  | 
| 
 | 
  3568  | 
       apply (fast intro: nth_mem[OF len])
  | 
| 
 | 
  3569  | 
      apply fast
  | 
| 
 | 
  3570  | 
     apply fast
  | 
| 
 | 
  3571  | 
    apply assumption
  | 
| 
 | 
  3572  | 
  done
  | 
| 
 | 
  3573  | 
  with len carr aa1carr aa2carr aa1fs
  | 
| 
 | 
  3574  | 
      have v2: "wfactors G (take i as' @ as'!i # drop (Suc i) as') (aa_1 \<otimes> (as'!i \<otimes> aa_2))"
  | 
| 
 | 
  3575  | 
    apply (intro wfactors_mult)
  | 
| 
 | 
  3576  | 
         apply fast
  | 
| 
 | 
  3577  | 
        apply (simp, (fast intro: nth_mem[OF len])?)+
  | 
| 
 | 
  3578  | 
  done
  | 
| 
 | 
  3579  | 
  | 
| 
 | 
  3580  | 
  from len
  | 
| 
 | 
  3581  | 
      have as': "as' = (take i as' @ as'!i # drop (Suc i) as')"
  | 
| 
 | 
  3582  | 
      by (simp add: drop_Suc_conv_tl)
  | 
| 
 | 
  3583  | 
  with carr
  | 
| 
 | 
  3584  | 
      have eer: "essentially_equal G (take i as' @ as'!i # drop (Suc i) as') as'"
  | 
| 
 | 
  3585  | 
      by simp
  | 
| 
 | 
  3586  | 
  | 
| 
 | 
  3587  | 
  with v2 afs' carr aa1carr aa2carr nth_mem[OF len]
  | 
| 
 | 
  3588  | 
      have "aa_1 \<otimes> (as'!i \<otimes> aa_2) \<sim> a"
  | 
| 
 | 
  3589  | 
    apply (intro ee_wfactorsD[of "take i as' @ as'!i # drop (Suc i) as'"  "as'"])
  | 
| 
 | 
  3590  | 
          apply fast+
  | 
| 
 | 
  3591  | 
        apply (simp, fast)
  | 
| 
 | 
  3592  | 
  done
  | 
| 
 | 
  3593  | 
  then
  | 
| 
 | 
  3594  | 
  have t1: "as'!i \<otimes> (aa_1 \<otimes> aa_2) \<sim> a"
  | 
| 
 | 
  3595  | 
    apply (simp add: m_assoc[symmetric])
  | 
| 
 | 
  3596  | 
    apply (simp add: m_comm)
  | 
| 
 | 
  3597  | 
  done
  | 
| 
 | 
  3598  | 
  from carr asiah
  | 
| 
 | 
  3599  | 
  have "ah \<otimes> (aa_1 \<otimes> aa_2) \<sim> as'!i \<otimes> (aa_1 \<otimes> aa_2)"
  | 
| 
 | 
  3600  | 
      apply (intro mult_cong_l)
  | 
| 
 | 
  3601  | 
      apply (fast intro: associated_sym, simp+)
  | 
| 
 | 
  3602  | 
  done
  | 
| 
 | 
  3603  | 
  also note t1
  | 
| 
 | 
  3604  | 
  finally
  | 
| 
 | 
  3605  | 
      have "ah \<otimes> (aa_1 \<otimes> aa_2) \<sim> a" by simp
  | 
| 
 | 
  3606  | 
  | 
| 
 | 
  3607  | 
  with carr aa1carr aa2carr a'carr nth_mem[OF len]
  | 
| 
 | 
  3608  | 
      have a': "aa_1 \<otimes> aa_2 \<sim> a'"
  | 
| 
 | 
  3609  | 
      by (simp add: a, fast intro: assoc_l_cancel[of ah _ a'])
  | 
| 
 | 
  3610  | 
  | 
| 
 | 
  3611  | 
  note v1
  | 
| 
 | 
  3612  | 
  also note a'
  | 
| 
 | 
  3613  | 
  finally have "wfactors G (take i as' @ drop (Suc i) as') a'" by simp
  | 
| 
 | 
  3614  | 
  | 
| 
 | 
  3615  | 
  from a'fs this carr
  | 
| 
 | 
  3616  | 
      have "essentially_equal G as (take i as' @ drop (Suc i) as')"
  | 
| 
 | 
  3617  | 
      by (intro ih[of a']) simp
  | 
| 
 | 
  3618  | 
  | 
| 
 | 
  3619  | 
  hence ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')"
  | 
| 
 | 
  3620  | 
    apply (elim essentially_equalE) apply (fastsimp intro: essentially_equalI)
  | 
| 
 | 
  3621  | 
  done
  | 
| 
 | 
  3622  | 
  | 
| 
 | 
  3623  | 
  from carr
  | 
| 
 | 
  3624  | 
  have ee2: "essentially_equal G (ah # take i as' @ drop (Suc i) as')
  | 
| 
 | 
  3625  | 
                                 (as' ! i # take i as' @ drop (Suc i) as')"
  | 
| 
 | 
  3626  | 
  proof (intro essentially_equalI)
  | 
| 
 | 
  3627  | 
    show "ah # take i as' @ drop (Suc i) as' <~~> ah # take i as' @ drop (Suc i) as'"
  | 
| 
 | 
  3628  | 
        by simp
  | 
| 
 | 
  3629  | 
  next show "ah # take i as' @ drop (Suc i) as' [\<sim>]
  | 
| 
 | 
  3630  | 
       as' ! i # take i as' @ drop (Suc i) as'"
  | 
| 
 | 
  3631  | 
    apply (simp add: list_all2_append)
  | 
| 
 | 
  3632  | 
    apply (simp add: asiah[symmetric] ahcarr asicarr)
  | 
| 
 | 
  3633  | 
    done
  | 
| 
 | 
  3634  | 
  qed
  | 
| 
 | 
  3635  | 
  | 
| 
 | 
  3636  | 
  note ee1
  | 
| 
 | 
  3637  | 
  also note ee2
  | 
| 
 | 
  3638  | 
  also have "essentially_equal G (as' ! i # take i as' @ drop (Suc i) as')
  | 
| 
 | 
  3639  | 
                                 (take i as' @ as' ! i # drop (Suc i) as')"
  | 
| 
 | 
  3640  | 
    apply (intro essentially_equalI)
  | 
| 
 | 
  3641  | 
    apply (subgoal_tac "as' ! i # take i as' @ drop (Suc i) as' <~~> 
  | 
| 
 | 
  3642  | 
                        take i as' @ as' ! i # drop (Suc i) as'")
  | 
| 
 | 
  3643  | 
apply simp
  | 
| 
 | 
  3644  | 
     apply (rule perm_append_Cons)
  | 
| 
 | 
  3645  | 
    apply simp
  | 
| 
 | 
  3646  | 
  done
  | 
| 
 | 
  3647  | 
  finally
  | 
| 
 | 
  3648  | 
      have "essentially_equal G (ah # as) 
  | 
| 
 | 
  3649  | 
                                (take i as' @ as' ! i # drop (Suc i) as')" by simp
  | 
| 
 | 
  3650  | 
  | 
| 
 | 
  3651  | 
  thus "essentially_equal G (ah # as) as'" by (subst as', assumption)
  | 
| 
 | 
  3652  | 
qed
  | 
| 
 | 
  3653  | 
  | 
| 
 | 
  3654  | 
lemma (in primeness_condition_monoid) wfactors_unique:
  | 
| 
 | 
  3655  | 
  assumes "wfactors G as a"  "wfactors G as' a"
  | 
| 
 | 
  3656  | 
    and "a \<in> carrier G"  "set as \<subseteq> carrier G"  "set as' \<subseteq> carrier G"
  | 
| 
 | 
  3657  | 
  shows "essentially_equal G as as'"
  | 
| 
 | 
  3658  | 
apply (rule wfactors_unique__hlp_induct[rule_format, of a])
  | 
| 
 | 
  3659  | 
apply (simp add: assms)
  | 
| 
 | 
  3660  | 
done
  | 
| 
 | 
  3661  | 
  | 
| 
 | 
  3662  | 
  | 
| 
 | 
  3663  | 
subsubsection {* Application to factorial monoids *}
 | 
| 
 | 
  3664  | 
  | 
| 
 | 
  3665  | 
text {* Number of factors for wellfoundedness *}
 | 
| 
 | 
  3666  | 
  | 
| 
 | 
  3667  | 
constdefs
  | 
| 
 | 
  3668  | 
  factorcount :: "_ \<Rightarrow> 'a \<Rightarrow> nat"
  | 
| 
 | 
  3669  | 
  "factorcount G a == THE c. (ALL as. set as \<subseteq> carrier G \<and> 
  | 
| 
 | 
  3670  | 
                                      wfactors G as a \<longrightarrow> c = length as)"
  | 
| 
 | 
  3671  | 
  | 
| 
 | 
  3672  | 
lemma (in monoid) ee_length:
  | 
| 
 | 
  3673  | 
  assumes ee: "essentially_equal G as bs"
  | 
| 
 | 
  3674  | 
  shows "length as = length bs"
  | 
| 
 | 
  3675  | 
apply (rule essentially_equalE[OF ee])
  | 
| 
 | 
  3676  | 
apply (subgoal_tac "length as = length fs1'")
  | 
| 
 | 
  3677  | 
 apply (simp add: list_all2_lengthD)
  | 
| 
 | 
  3678  | 
apply (simp add: perm_length)
  | 
| 
 | 
  3679  | 
done
  | 
| 
 | 
  3680  | 
  | 
| 
 | 
  3681  | 
lemma (in factorial_monoid) factorcount_exists:
  | 
| 
 | 
  3682  | 
  assumes carr[simp]: "a \<in> carrier G"
  | 
| 
 | 
  3683  | 
  shows "EX c. ALL as. set as \<subseteq> carrier G \<and> wfactors G as a \<longrightarrow> c = length as"
  | 
| 
 | 
  3684  | 
proof -
  | 
| 
 | 
  3685  | 
  have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a" by (intro wfactors_exist, simp)
  | 
| 
 | 
  3686  | 
  from this obtain as
  | 
| 
 | 
  3687  | 
      where ascarr[simp]: "set as \<subseteq> carrier G"
  | 
| 
 | 
  3688  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  3689  | 
      by (auto simp del: carr)
  | 
| 
 | 
  3690  | 
  | 
| 
 | 
  3691  | 
  have "ALL as'. set as' \<subseteq> carrier G \<and> wfactors G as' a \<longrightarrow> length as = length as'"
  | 
| 
 | 
  3692  | 
  proof clarify
  | 
| 
 | 
  3693  | 
    fix as'
  | 
| 
 | 
  3694  | 
    assume [simp]: "set as' \<subseteq> carrier G"
  | 
| 
 | 
  3695  | 
      and bfs: "wfactors G as' a"
  | 
| 
 | 
  3696  | 
    from afs bfs
  | 
| 
 | 
  3697  | 
        have "essentially_equal G as as'"
  | 
| 
 | 
  3698  | 
        by (intro ee_wfactorsI[of a a as as'], simp+)
  | 
| 
 | 
  3699  | 
    thus "length as = length as'" by (rule ee_length)
  | 
| 
 | 
  3700  | 
  qed
  | 
| 
 | 
  3701  | 
  thus "EX c. ALL as'. set as' \<subseteq> carrier G \<and> wfactors G as' a \<longrightarrow> c = length as'" ..
  | 
| 
 | 
  3702  | 
qed
  | 
| 
 | 
  3703  | 
  | 
| 
 | 
  3704  | 
lemma (in factorial_monoid) factorcount_unique:
  | 
| 
 | 
  3705  | 
  assumes afs: "wfactors G as a"
  | 
| 
 | 
  3706  | 
    and acarr[simp]: "a \<in> carrier G" and ascarr[simp]: "set as \<subseteq> carrier G"
  | 
| 
 | 
  3707  | 
  shows "factorcount G a = length as"
  | 
| 
 | 
  3708  | 
proof -
  | 
| 
 | 
  3709  | 
  have "EX ac. ALL as. set as \<subseteq> carrier G \<and> wfactors G as a \<longrightarrow> ac = length as" by (rule factorcount_exists, simp)
  | 
| 
 | 
  3710  | 
  from this obtain ac where
  | 
| 
 | 
  3711  | 
      alen: "ALL as. set as \<subseteq> carrier G \<and> wfactors G as a \<longrightarrow> ac = length as"
  | 
| 
 | 
  3712  | 
      by auto
  | 
| 
 | 
  3713  | 
  have ac: "ac = factorcount G a"
  | 
| 
 | 
  3714  | 
    apply (simp add: factorcount_def)
  | 
| 
 | 
  3715  | 
    apply (rule theI2)
  | 
| 
 | 
  3716  | 
      apply (rule alen)
  | 
| 
 | 
  3717  | 
     apply (elim allE[of _ "as"], rule allE[OF alen, of "as"], simp add: ascarr afs)
  | 
| 
 | 
  3718  | 
    apply (elim allE[of _ "as"], rule allE[OF alen, of "as"], simp add: ascarr afs)
  | 
| 
 | 
  3719  | 
  done
  | 
| 
 | 
  3720  | 
  | 
| 
 | 
  3721  | 
  from ascarr afs have "ac = length as" by (iprover intro: alen[rule_format])
  | 
| 
 | 
  3722  | 
  with ac show ?thesis by simp
  | 
| 
 | 
  3723  | 
qed
  | 
| 
 | 
  3724  | 
  | 
| 
 | 
  3725  | 
lemma (in factorial_monoid) divides_fcount:
  | 
| 
 | 
  3726  | 
  assumes dvd: "a divides b"
  | 
| 
 | 
  3727  | 
    and acarr: "a \<in> carrier G" and bcarr:"b \<in> carrier G"
  | 
| 
 | 
  3728  | 
  shows "factorcount G a <= factorcount G b"
  | 
| 
 | 
  3729  | 
apply (rule dividesE[OF dvd])
  | 
| 
 | 
  3730  | 
proof -
  | 
| 
 | 
  3731  | 
  fix c
  | 
| 
 | 
  3732  | 
  from assms
  | 
| 
 | 
  3733  | 
      have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a" by fast
  | 
| 
 | 
  3734  | 
  from this obtain as
  | 
| 
 | 
  3735  | 
      where ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  3736  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  3737  | 
      by auto
  | 
| 
 | 
  3738  | 
  with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique)
  | 
| 
 | 
  3739  | 
  | 
| 
 | 
  3740  | 
  assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
  3741  | 
  hence "\<exists>cs. set cs \<subseteq> carrier G \<and> wfactors G cs c" by fast
  | 
| 
 | 
  3742  | 
  from this obtain cs
  | 
| 
 | 
  3743  | 
      where cscarr: "set cs \<subseteq> carrier G"
  | 
| 
 | 
  3744  | 
      and cfs: "wfactors G cs c"
  | 
| 
 | 
  3745  | 
      by auto
  | 
| 
 | 
  3746  | 
  | 
| 
 | 
  3747  | 
  note [simp] = acarr bcarr ccarr ascarr cscarr
  | 
| 
 | 
  3748  | 
  | 
| 
 | 
  3749  | 
  assume b: "b = a \<otimes> c"
  | 
| 
 | 
  3750  | 
  from afs cfs
  | 
| 
 | 
  3751  | 
      have "wfactors G (as@cs) (a \<otimes> c)" by (intro wfactors_mult, simp+)
  | 
| 
 | 
  3752  | 
  with b have "wfactors G (as@cs) b" by simp
  | 
| 
 | 
  3753  | 
  hence "factorcount G b = length (as@cs)" by (intro factorcount_unique, simp+)
  | 
| 
 | 
  3754  | 
  hence "factorcount G b = length as + length cs" by simp
  | 
| 
 | 
  3755  | 
  with fca show ?thesis by simp
  | 
| 
 | 
  3756  | 
qed
  | 
| 
 | 
  3757  | 
  | 
| 
 | 
  3758  | 
lemma (in factorial_monoid) associated_fcount:
  | 
| 
 | 
  3759  | 
  assumes acarr: "a \<in> carrier G" and bcarr:"b \<in> carrier G"
  | 
| 
 | 
  3760  | 
    and asc: "a \<sim> b"
  | 
| 
 | 
  3761  | 
  shows "factorcount G a = factorcount G b"
  | 
| 
 | 
  3762  | 
apply (rule associatedE[OF asc])
  | 
| 
 | 
  3763  | 
apply (drule divides_fcount[OF _ acarr bcarr])
  | 
| 
 | 
  3764  | 
apply (drule divides_fcount[OF _ bcarr acarr])
  | 
| 
 | 
  3765  | 
apply simp
  | 
| 
 | 
  3766  | 
done
  | 
| 
 | 
  3767  | 
  | 
| 
 | 
  3768  | 
lemma (in factorial_monoid) properfactor_fcount:
  | 
| 
 | 
  3769  | 
  assumes acarr: "a \<in> carrier G" and bcarr:"b \<in> carrier G"
  | 
| 
 | 
  3770  | 
    and pf: "properfactor G a b"
  | 
| 
 | 
  3771  | 
  shows "factorcount G a < factorcount G b"
  | 
| 
 | 
  3772  | 
apply (rule properfactorE[OF pf], elim dividesE)
  | 
| 
 | 
  3773  | 
proof -
  | 
| 
 | 
  3774  | 
  fix c
  | 
| 
 | 
  3775  | 
  from assms
  | 
| 
 | 
  3776  | 
  have "\<exists>as. set as \<subseteq> carrier G \<and> wfactors G as a" by fast
  | 
| 
 | 
  3777  | 
  from this obtain as
  | 
| 
 | 
  3778  | 
      where ascarr: "set as \<subseteq> carrier G"
  | 
| 
 | 
  3779  | 
      and afs: "wfactors G as a"
  | 
| 
 | 
  3780  | 
      by auto
  | 
| 
 | 
  3781  | 
  with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique)
  | 
| 
 | 
  3782  | 
  | 
| 
 | 
  3783  | 
  assume ccarr: "c \<in> carrier G"
  | 
| 
 | 
  3784  | 
  hence "\<exists>cs. set cs \<subseteq> carrier G \<and> wfactors G cs c" by fast
  | 
| 
 | 
  3785  | 
  from this obtain cs
  | 
| 
 | 
  3786  | 
      where cscarr: "set cs \<subseteq> carrier G"
  | 
| 
 | 
  3787  | 
      and cfs: "wfactors G cs c"
  | 
| 
 | 
  3788  | 
      by auto
  | 
| 
 | 
  3789  | 
  | 
| 
 | 
  3790  | 
  assume b: "b = a \<otimes> c"
  | 
| 
 | 
  3791  | 
  | 
| 
 | 
  3792  | 
  have "wfactors G (as@cs) (a \<otimes> c)" by (rule wfactors_mult) fact+
  | 
| 
 | 
  3793  | 
  with b
  | 
| 
 | 
  3794  | 
      have "wfactors G (as@cs) b" by simp
  | 
| 
 | 
  3795  | 
  with ascarr cscarr bcarr
  | 
| 
 | 
  3796  | 
      have "factorcount G b = length (as@cs)" by (simp add: factorcount_unique)
  | 
| 
 | 
  3797  | 
  hence fcb: "factorcount G b = length as + length cs" by simp
  | 
| 
 | 
  3798  | 
  | 
| 
 | 
  3799  | 
  assume nbdvda: "\<not> b divides a"
  | 
| 
 | 
  3800  | 
  have "c \<notin> Units G"
  | 
| 
 | 
  3801  | 
  proof (rule ccontr, simp)
  | 
| 
 | 
  3802  | 
    assume cunit:"c \<in> Units G"
  | 
| 
 | 
  3803  | 
  | 
| 
 | 
  3804  | 
    have "b \<otimes> inv c = a \<otimes> c \<otimes> inv c" by (simp add: b)
  | 
| 
 | 
  3805  | 
    also with ccarr acarr cunit
  | 
| 
 | 
  3806  | 
        have "\<dots> = a \<otimes> (c \<otimes> inv c)" by (fast intro: m_assoc)
  | 
| 
 | 
  3807  | 
    also with ccarr cunit
  | 
| 
 | 
  3808  | 
        have "\<dots> = a \<otimes> \<one>" by (simp add: Units_r_inv)
  | 
| 
 | 
  3809  | 
    also with acarr
  | 
| 
 | 
  3810  | 
        have "\<dots> = a" by simp
  | 
| 
 | 
  3811  | 
    finally have "a = b \<otimes> inv c" by simp
  | 
| 
 | 
  3812  | 
    with ccarr cunit
  | 
| 
 | 
  3813  | 
    have "b divides a" by (fast intro: dividesI[of "inv c"])
  | 
| 
 | 
  3814  | 
    with nbdvda show False by simp
  | 
| 
 | 
  3815  | 
  qed
  | 
| 
 | 
  3816  | 
  | 
| 
 | 
  3817  | 
  with cfs have "length cs > 0"
  | 
| 
 | 
  3818  | 
  apply -
  | 
| 
 | 
  3819  | 
  apply (rule ccontr, simp)
  | 
| 
 | 
  3820  | 
  proof -
  | 
| 
 | 
  3821  | 
    assume "wfactors G [] c"
  | 
| 
 | 
  3822  | 
    hence "\<one> \<sim> c" by (elim wfactorsE, simp)
  | 
| 
 | 
  3823  | 
    with ccarr
  | 
| 
 | 
  3824  | 
        have cunit: "c \<in> Units G" by (intro assoc_unit_r[of "\<one>" "c"], simp+)
  | 
| 
 | 
  3825  | 
    assume "c \<notin> Units G"
  | 
| 
 | 
  3826  | 
    with cunit show "False" by simp
  | 
| 
 | 
  3827  | 
  qed
  | 
| 
 | 
  3828  | 
  | 
| 
 | 
  3829  | 
  with fca fcb show ?thesis by simp
  | 
| 
 | 
  3830  | 
qed
  | 
| 
 | 
  3831  | 
  | 
| 
 | 
  3832  | 
interpretation factorial_monoid \<subseteq> divisor_chain_condition_monoid
  | 
| 
 | 
  3833  | 
apply unfold_locales
  | 
| 
 | 
  3834  | 
apply (rule wfUNIVI)
  | 
| 
 | 
  3835  | 
apply (rule measure_induct[of "factorcount G"])
  | 
| 
 | 
  3836  | 
apply simp (* slow *) (*
  | 
| 
 | 
  3837  | 
  [1]Applying congruence rule:
  | 
| 
 | 
  3838  | 
  \<lbrakk>factorcount G y < factorcount G xa \<equiv> ?P'; ?P' \<Longrightarrow> P y \<equiv> ?Q'\<rbrakk> \<Longrightarrow> factorcount G y < factorcount G xa \<longrightarrow> P y \<equiv> ?P' \<longrightarrow> ?Q'
  | 
| 
 | 
  3839  | 
  | 
| 
 | 
  3840  | 
  trace_simp_depth_limit exceeded!
  | 
| 
 | 
  3841  | 
*)
  | 
| 
 | 
  3842  | 
proof -
  | 
| 
 | 
  3843  | 
  fix P x
  | 
| 
 | 
  3844  | 
  assume r1[rule_format]:
  | 
| 
 | 
  3845  | 
         "\<forall>y. (\<forall>z. z \<in> carrier G \<and> y \<in> carrier G \<and> properfactor G z y \<longrightarrow> P z) \<longrightarrow> P y"
  | 
| 
 | 
  3846  | 
    and r2[rule_format]: "\<forall>y. factorcount G y < factorcount G x \<longrightarrow> P y"
  | 
| 
 | 
  3847  | 
  show "P x"
  | 
| 
 | 
  3848  | 
    apply (rule r1)
  | 
| 
 | 
  3849  | 
    apply (rule r2)
  | 
| 
 | 
  3850  | 
    apply (rule properfactor_fcount, simp+)
  | 
| 
 | 
  3851  | 
  done
  | 
| 
 | 
  3852  | 
qed
  | 
| 
 | 
  3853  | 
  | 
| 
 | 
  3854  | 
interpretation factorial_monoid \<subseteq> primeness_condition_monoid
  | 
| 
 | 
  3855  | 
  by (unfold_locales, rule irreducible_is_prime)
  | 
| 
 | 
  3856  | 
  | 
| 
 | 
  3857  | 
  | 
| 
 | 
  3858  | 
lemma (in factorial_monoid) primeness_condition:
  | 
| 
 | 
  3859  | 
  shows "primeness_condition_monoid G"
  | 
| 
 | 
  3860  | 
by unfold_locales
  | 
| 
 | 
  3861  | 
  | 
| 
 | 
  3862  | 
lemma (in factorial_monoid) gcd_condition [simp]:
  | 
| 
 | 
  3863  | 
  shows "gcd_condition_monoid G"
  | 
| 
 | 
  3864  | 
by (unfold_locales, rule gcdof_exists)
  | 
| 
 | 
  3865  | 
  | 
| 
 | 
  3866  | 
interpretation factorial_monoid \<subseteq> gcd_condition_monoid
  | 
| 
 | 
  3867  | 
  by (unfold_locales, rule gcdof_exists)
  | 
| 
 | 
  3868  | 
  | 
| 
 | 
  3869  | 
lemma (in factorial_monoid) division_glattice [simp]:
  | 
| 
 | 
  3870  | 
  shows "glattice (division_rel G)"
  | 
| 
 | 
  3871  | 
proof -
  | 
| 
 | 
  3872  | 
  interpret glower_semilattice ["division_rel G"] by simp
  | 
| 
 | 
  3873  | 
  | 
| 
 | 
  3874  | 
  show "glattice (division_rel G)"
  | 
| 
 | 
  3875  | 
  apply (unfold_locales, simp_all)
  | 
| 
 | 
  3876  | 
  proof -
  | 
| 
 | 
  3877  | 
    fix x y
  | 
| 
 | 
  3878  | 
    assume carr: "x \<in> carrier G"  "y \<in> carrier G"
  | 
| 
 | 
  3879  | 
  | 
| 
 | 
  3880  | 
    hence "\<exists>z. z \<in> carrier G \<and> z lcmof x y" by (rule lcmof_exists)
  | 
| 
 | 
  3881  | 
    from this obtain z
  | 
| 
 | 
  3882  | 
        where zcarr: "z \<in> carrier G"
  | 
| 
 | 
  3883  | 
        and isgcd: "z lcmof x y"
  | 
| 
 | 
  3884  | 
        by auto
  | 
| 
 | 
  3885  | 
    with carr
  | 
| 
 | 
  3886  | 
    have "gleast (division_rel G) z (Upper (division_rel G) {x, y})"
 | 
| 
 | 
  3887  | 
        by (simp add: lcmof_gleastUpper[symmetric])
  | 
| 
 | 
  3888  | 
    thus "\<exists>z. gleast (division_rel G) z (Upper (division_rel G) {x, y})" by fast
 | 
| 
 | 
  3889  | 
  qed
  | 
| 
 | 
  3890  | 
qed
  | 
| 
 | 
  3891  | 
  | 
| 
 | 
  3892  | 
  | 
| 
 | 
  3893  | 
subsection {* Factoriality theorems *}
 | 
| 
 | 
  3894  | 
  | 
| 
 | 
  3895  | 
theorem factorial_condition_one: (* Jacobson theorem 2.21 *)
  | 
| 
 | 
  3896  | 
  shows "(divisor_chain_condition_monoid G \<and> primeness_condition_monoid G) = 
  | 
| 
 | 
  3897  | 
         factorial_monoid G"
  | 
| 
 | 
  3898  | 
apply rule
  | 
| 
 | 
  3899  | 
proof clarify
  | 
| 
 | 
  3900  | 
  assume dcc: "divisor_chain_condition_monoid G"
  | 
| 
 | 
  3901  | 
     and pc: "primeness_condition_monoid G"
  | 
| 
 | 
  3902  | 
  interpret divisor_chain_condition_monoid ["G"] by (rule dcc)
  | 
| 
 | 
  3903  | 
  interpret primeness_condition_monoid ["G"] by (rule pc)
  | 
| 
 | 
  3904  | 
  | 
| 
 | 
  3905  | 
  show "factorial_monoid G"
  | 
| 
 | 
  3906  | 
      by (fast intro: factorial_monoidI wfactors_exist wfactors_unique)
  | 
| 
 | 
  3907  | 
next
  | 
| 
 | 
  3908  | 
  assume fm: "factorial_monoid G"
  | 
| 
 | 
  3909  | 
  interpret factorial_monoid ["G"] by (rule fm)
  | 
| 
 | 
  3910  | 
  show "divisor_chain_condition_monoid G \<and> primeness_condition_monoid G"
  | 
| 
 | 
  3911  | 
      by rule unfold_locales
  | 
| 
 | 
  3912  | 
qed
  | 
| 
 | 
  3913  | 
  | 
| 
 | 
  3914  | 
theorem factorial_condition_two: (* Jacobson theorem 2.22 *)
  | 
| 
 | 
  3915  | 
  shows "(divisor_chain_condition_monoid G \<and> gcd_condition_monoid G) = factorial_monoid G"
  | 
| 
 | 
  3916  | 
apply rule
  | 
| 
 | 
  3917  | 
proof clarify
  | 
| 
 | 
  3918  | 
    assume dcc: "divisor_chain_condition_monoid G"
  | 
| 
 | 
  3919  | 
     and gc: "gcd_condition_monoid G"
  | 
| 
 | 
  3920  | 
  interpret divisor_chain_condition_monoid ["G"] by (rule dcc)
  | 
| 
 | 
  3921  | 
  interpret gcd_condition_monoid ["G"] by (rule gc)
  | 
| 
 | 
  3922  | 
  show "factorial_monoid G"
  | 
| 
 | 
  3923  | 
      by (simp add: factorial_condition_one[symmetric], rule, unfold_locales)
  | 
| 
 | 
  3924  | 
next
  | 
| 
 | 
  3925  | 
  assume fm: "factorial_monoid G"
  | 
| 
 | 
  3926  | 
  interpret factorial_monoid ["G"] by (rule fm)
  | 
| 
 | 
  3927  | 
  show "divisor_chain_condition_monoid G \<and> gcd_condition_monoid G"
  | 
| 
 | 
  3928  | 
      by rule unfold_locales
  | 
| 
 | 
  3929  | 
qed
  | 
| 
 | 
  3930  | 
  | 
| 
 | 
  3931  | 
end
  |