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