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