src/HOL/Decision_Procs/Algebra_Aux.thy
author nipkow
Fri, 05 Jan 2018 18:41:42 +0100
changeset 67341 df79ef3b3a41
parent 67123 3fe40ff1b921
child 67397 12b0c11e3d20
permissions -rw-r--r--
Renamed (^) to [^] in preparation of the move from "op X" to (X)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Decision_Procs/Algebra_Aux.thy
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     3
*)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     4
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     5
section \<open>Things that can be added to the Algebra library\<close>
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     6
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     7
theory Algebra_Aux
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
     8
  imports "HOL-Algebra.Ring"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
     9
begin
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    10
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    11
definition of_natural :: "('a, 'm) ring_scheme \<Rightarrow> nat \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>")
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    12
  where "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> = (op \<oplus>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub> ^^ n) \<zero>\<^bsub>R\<^esub>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    13
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    14
definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<index>")
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    15
  where "\<guillemotleft>i\<guillemotright>\<^bsub>R\<^esub> = (if 0 \<le> i then \<guillemotleft>nat i\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> else \<ominus>\<^bsub>R\<^esub> \<guillemotleft>nat (- i)\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub>)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    16
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    17
context ring
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    18
begin
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    19
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    20
lemma of_nat_0 [simp]: "\<guillemotleft>0\<guillemotright>\<^sub>\<nat> = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    21
  by (simp add: of_natural_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    22
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    23
lemma of_nat_Suc [simp]: "\<guillemotleft>Suc n\<guillemotright>\<^sub>\<nat> = \<one> \<oplus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    24
  by (simp add: of_natural_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    25
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    26
lemma of_int_0 [simp]: "\<guillemotleft>0\<guillemotright> = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    27
  by (simp add: of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    28
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    29
lemma of_nat_closed [simp]: "\<guillemotleft>n\<guillemotright>\<^sub>\<nat> \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    30
  by (induct n) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    31
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    32
lemma of_int_closed [simp]: "\<guillemotleft>i\<guillemotright> \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    33
  by (simp add: of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    34
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    35
lemma of_int_minus [simp]: "\<guillemotleft>- i\<guillemotright> = \<ominus> \<guillemotleft>i\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    36
  by (simp add: of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    37
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    38
lemma of_nat_add [simp]: "\<guillemotleft>m + n\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<oplus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    39
  by (induct m) (simp_all add: a_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    40
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    41
lemma of_nat_diff [simp]: "n \<le> m \<Longrightarrow> \<guillemotleft>m - n\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>n\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    42
proof (induct m arbitrary: n)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    43
  case 0
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    44
  then show ?case by (simp add: minus_eq)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    45
next
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    46
  case Suc': (Suc m)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    47
  show ?case
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    48
  proof (cases n)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    49
    case 0
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    50
    then show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    51
      by (simp add: minus_eq)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    52
  next
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    53
    case (Suc k)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    54
    with Suc' have "\<guillemotleft>Suc m - Suc k\<guillemotright>\<^sub>\<nat> = \<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    55
    also have "\<dots> = \<one> \<oplus> \<ominus> \<one> \<oplus> (\<guillemotleft>m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>k\<guillemotright>\<^sub>\<nat>)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    56
      by (simp add: r_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    57
    also have "\<dots> = \<guillemotleft>Suc m\<guillemotright>\<^sub>\<nat> \<ominus> \<guillemotleft>Suc k\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    58
      by (simp add: minus_eq minus_add a_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    59
    finally show ?thesis using Suc by simp
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    60
  qed
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    61
qed
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    62
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    63
lemma of_int_add [simp]: "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<oplus> \<guillemotleft>j\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    64
proof (cases "0 \<le> i")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    65
  case True
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    66
  show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    67
  proof (cases "0 \<le> j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    68
    case True
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    69
    with \<open>0 \<le> i\<close> show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
    70
      by (simp add: of_integer_def nat_add_distrib)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    71
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    72
    case False
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    73
    show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    74
    proof (cases "0 \<le> i + j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    75
      case True
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    76
      then have "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>nat (i - (- j))\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    77
        by (simp add: of_integer_def)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
    78
      also from True \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    79
      have "nat (i - (- j)) = nat i - nat (- j)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    80
        by (simp add: nat_diff_distrib)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
    81
      finally show ?thesis using True \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    82
        by (simp add: minus_eq of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    83
    next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    84
      case False
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    85
      then have "\<guillemotleft>i + j\<guillemotright> = \<ominus> \<guillemotleft>nat (- j - i)\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    86
        by (simp add: of_integer_def) (simp only: diff_conv_add_uminus add_ac)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
    87
      also from False \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    88
      have "nat (- j - i) = nat (- j) - nat i"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    89
        by (simp add: nat_diff_distrib)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
    90
      finally show ?thesis using False \<open>0 \<le> i\<close> \<open>\<not> 0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    91
        by (simp add: minus_eq minus_add a_ac of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    92
    qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    93
  qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    94
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    95
  case False
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    96
  show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    97
  proof (cases "0 \<le> j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    98
    case True
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
    99
    show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   100
    proof (cases "0 \<le> i + j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   101
      case True
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   102
      then have "\<guillemotleft>i + j\<guillemotright> = \<guillemotleft>nat (j - (- i))\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   103
        by (simp add: of_integer_def add_ac)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   104
      also from True \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   105
      have "nat (j - (- i)) = nat j - nat (- i)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   106
        by (simp add: nat_diff_distrib)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   107
      finally show ?thesis using True \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   108
        by (simp add: minus_eq minus_add a_ac of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   109
    next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   110
      case False
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   111
      then have "\<guillemotleft>i + j\<guillemotright> = \<ominus> \<guillemotleft>nat (- i - j)\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   112
        by (simp add: of_integer_def)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   113
      also from False \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   114
      have "nat (- i - j) = nat (- i) - nat j"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   115
        by (simp add: nat_diff_distrib)
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   116
      finally show ?thesis using False \<open>\<not> 0 \<le> i\<close> \<open>0 \<le> j\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   117
        by (simp add: minus_eq minus_add of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   118
    qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   119
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   120
    case False
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   121
    with \<open>\<not> 0 \<le> i\<close> show ?thesis
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   122
      by (simp add: of_integer_def nat_add_distrib minus_add diff_conv_add_uminus
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   123
          del: add_uminus_conv_diff uminus_add_conv_diff)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   124
  qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   125
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   126
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   127
lemma of_int_diff [simp]: "\<guillemotleft>i - j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<ominus> \<guillemotleft>j\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   128
  by (simp only: diff_conv_add_uminus of_int_add) (simp add: minus_eq)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   129
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   130
lemma of_nat_mult [simp]: "\<guillemotleft>i * j\<guillemotright>\<^sub>\<nat> = \<guillemotleft>i\<guillemotright>\<^sub>\<nat> \<otimes> \<guillemotleft>j\<guillemotright>\<^sub>\<nat>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   131
  by (induct i) (simp_all add: l_distr)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   132
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   133
lemma of_int_mult [simp]: "\<guillemotleft>i * j\<guillemotright> = \<guillemotleft>i\<guillemotright> \<otimes> \<guillemotleft>j\<guillemotright>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   134
proof (cases "0 \<le> i")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   135
  case True
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   136
  show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   137
  proof (cases "0 \<le> j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   138
    case True
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   139
    with \<open>0 \<le> i\<close> show ?thesis
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   140
      by (simp add: of_integer_def nat_mult_distrib zero_le_mult_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   141
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   142
    case False
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   143
    with \<open>0 \<le> i\<close> show ?thesis
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   144
      by (simp add: of_integer_def zero_le_mult_iff
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   145
        minus_mult_right nat_mult_distrib r_minus
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   146
        del: minus_mult_right [symmetric])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   147
  qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   148
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   149
  case False
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   150
  show ?thesis
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   151
  proof (cases "0 \<le> j")
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   152
    case True
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   153
    with \<open>\<not> 0 \<le> i\<close> show ?thesis
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   154
      by (simp add: of_integer_def zero_le_mult_iff
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   155
        minus_mult_left nat_mult_distrib l_minus
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   156
        del: minus_mult_left [symmetric])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   157
  next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   158
    case False
64998
d51478d6aae4 isabelle update_cartouches
haftmann
parents: 64962
diff changeset
   159
    with \<open>\<not> 0 \<le> i\<close> show ?thesis
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   160
      by (simp add: of_integer_def zero_le_mult_iff
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   161
        minus_mult_minus [of i j, symmetric] nat_mult_distrib
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   162
        l_minus r_minus
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   163
        del: minus_mult_minus
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   164
        minus_mult_left [symmetric] minus_mult_right [symmetric])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   165
  qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   166
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   167
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   168
lemma of_int_1 [simp]: "\<guillemotleft>1\<guillemotright> = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   169
  by (simp add: of_integer_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   170
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   171
lemma of_int_2: "\<guillemotleft>2\<guillemotright> = \<one> \<oplus> \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   172
  by (simp add: of_integer_def numeral_2_eq_2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   173
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   174
lemma minus_0_r [simp]: "x \<in> carrier R \<Longrightarrow> x \<ominus> \<zero> = x"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   175
  by (simp add: minus_eq)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   176
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   177
lemma minus_0_l [simp]: "x \<in> carrier R \<Longrightarrow> \<zero> \<ominus> x = \<ominus> x"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   178
  by (simp add: minus_eq)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   179
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   180
lemma eq_diff0:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   181
  assumes "x \<in> carrier R" "y \<in> carrier R"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   182
  shows "x \<ominus> y = \<zero> \<longleftrightarrow> x = y"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   183
proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   184
  assume "x \<ominus> y = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   185
  with assms have "x \<oplus> (\<ominus> y \<oplus> y) = y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   186
    by (simp add: minus_eq a_assoc [symmetric])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   187
  with assms show "x = y" by (simp add: l_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   188
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   189
  assume "x = y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   190
  with assms show "x \<ominus> y = \<zero>" by (simp add: minus_eq r_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   191
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   192
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   193
lemma power2_eq_square: "x \<in> carrier R \<Longrightarrow> x [^] (2::nat) = x \<otimes> x"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   194
  by (simp add: numeral_eq_Suc)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   195
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   196
lemma eq_neg_iff_add_eq_0:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   197
  assumes "x \<in> carrier R" "y \<in> carrier R"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   198
  shows "x = \<ominus> y \<longleftrightarrow> x \<oplus> y = \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   199
proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   200
  assume "x = \<ominus> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   201
  with assms show "x \<oplus> y = \<zero>" by (simp add: l_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   202
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   203
  assume "x \<oplus> y = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   204
  with assms have "x \<oplus> (y \<oplus> \<ominus> y) = \<zero> \<oplus> \<ominus> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   205
    by (simp add: a_assoc [symmetric])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   206
  with assms show "x = \<ominus> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   207
    by (simp add: r_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   208
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   209
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   210
lemma neg_equal_iff_equal:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   211
  assumes x: "x \<in> carrier R" and y: "y \<in> carrier R"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   212
  shows "\<ominus> x = \<ominus> y \<longleftrightarrow> x = y"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   213
proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   214
  assume "\<ominus> x = \<ominus> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   215
  then have "\<ominus> (\<ominus> x) = \<ominus> (\<ominus> y)" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   216
  with x y show "x = y" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   217
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   218
  assume "x = y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   219
  then show "\<ominus> x = \<ominus> y" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   220
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   221
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   222
lemma neg_equal_swap:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   223
  assumes x: "x \<in> carrier R" and y: "y \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   224
  shows "(\<ominus> x = y) = (x = \<ominus> y)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   225
  using assms neg_equal_iff_equal [of x "\<ominus> y"]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   226
  by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   227
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   228
lemma mult2: "x \<in> carrier R \<Longrightarrow> x \<oplus> x = \<guillemotleft>2\<guillemotright> \<otimes> x"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   229
  by (simp add: of_int_2 l_distr)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   230
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   231
end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   232
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   233
lemma (in cring) of_int_power [simp]: "\<guillemotleft>i ^ n\<guillemotright> = \<guillemotleft>i\<guillemotright> [^] n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   234
  by (induct n) (simp_all add: m_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   235
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   236
definition cring_class_ops :: "'a::comm_ring_1 ring"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   237
  where "cring_class_ops \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   238
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   239
lemma cring_class: "cring cring_class_ops"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   240
  apply unfold_locales
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   241
  apply (auto simp add: cring_class_ops_def ring_distribs Units_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   242
  apply (rule_tac x="- x" in exI)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   243
  apply simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   244
  done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   245
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   246
lemma carrier_class: "x \<in> carrier cring_class_ops"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   247
  by (simp add: cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   248
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   249
lemma zero_class: "\<zero>\<^bsub>cring_class_ops\<^esub> = 0"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   250
  by (simp add: cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   251
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   252
lemma one_class: "\<one>\<^bsub>cring_class_ops\<^esub> = 1"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   253
  by (simp add: cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   254
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   255
lemma plus_class: "x \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   256
  by (simp add: cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   257
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   258
lemma times_class: "x \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   259
  by (simp add: cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   260
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   261
lemma uminus_class: "\<ominus>\<^bsub>cring_class_ops\<^esub> x = - x"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   262
  apply (simp add: a_inv_def m_inv_def cring_class_ops_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   263
  apply (rule the_equality)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   264
  apply (simp_all add: eq_neg_iff_add_eq_0)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   265
  done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   266
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   267
lemma minus_class: "x \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   268
  by (simp add: a_minus_def carrier_class plus_class uminus_class)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   269
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   270
lemma power_class: "x [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   271
  by (induct n) (simp_all add: one_class times_class
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   272
    monoid.nat_pow_0 [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]]
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   273
    monoid.nat_pow_Suc [OF comm_monoid.axioms(1) [OF cring.axioms(2) [OF cring_class]]])
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   274
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   275
lemma of_nat_class: "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   276
  by (induct n) (simp_all add: cring_class_ops_def of_natural_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   277
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   278
lemma of_int_class: "\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub> = of_int i"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   279
  by (simp add: of_integer_def of_nat_class uminus_class)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   280
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   281
lemmas class_simps = zero_class one_class plus_class minus_class uminus_class
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   282
  times_class power_class of_nat_class of_int_class carrier_class
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   283
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   284
interpretation cring_class: cring "cring_class_ops::'a::comm_ring_1 ring"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   285
  rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   286
    and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   287
    and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   288
    and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   289
    and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   290
    and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   291
    and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   292
    and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   293
    and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   294
    and "(Int.of_int (numeral m)::'a) = numeral m"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   295
  by (simp_all add: cring_class class_simps)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   296
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   297
lemma (in domain) nat_pow_eq_0_iff [simp]:
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   298
  "a \<in> carrier R \<Longrightarrow> (a [^] (n::nat) = \<zero>) = (a = \<zero> \<and> n \<noteq> 0)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   299
  by (induct n) (auto simp add: integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   300
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   301
lemma (in domain) square_eq_iff:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   302
  assumes "x \<in> carrier R" "y \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   303
  shows "(x \<otimes> x = y \<otimes> y) = (x = y \<or> x = \<ominus> y)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   304
proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   305
  assume "x \<otimes> x = y \<otimes> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   306
  with assms have "(x \<ominus> y) \<otimes> (x \<oplus> y) = x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) \<oplus> (y \<otimes> y \<oplus> \<ominus> (y \<otimes> y))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   307
    by (simp add: r_distr l_distr minus_eq r_minus m_comm a_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   308
  with assms show "x = y \<or> x = \<ominus> y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   309
    by (simp add: integral_iff eq_neg_iff_add_eq_0 eq_diff0 r_neg)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   310
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   311
  assume "x = y \<or> x = \<ominus> y"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   312
  with assms show "x \<otimes> x = y \<otimes> y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   313
    by (auto simp add: l_minus r_minus)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   314
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   315
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   316
definition m_div :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oslash>\<index>" 70)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   317
  where "x \<oslash>\<^bsub>G\<^esub> y = (if y = \<zero>\<^bsub>G\<^esub> then \<zero>\<^bsub>G\<^esub> else x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   318
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   319
context field
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   320
begin
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   321
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   322
lemma inv_closed [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> inv x \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   323
  by (simp add: field_Units)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   324
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   325
lemma l_inv [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> inv x \<otimes> x = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   326
  by (simp add: field_Units)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   327
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   328
lemma r_inv [simp]: "x \<in> carrier R \<Longrightarrow> x \<noteq> \<zero> \<Longrightarrow> x \<otimes> inv x = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   329
  by (simp add: field_Units)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   330
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   331
lemma inverse_unique:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   332
  assumes a: "a \<in> carrier R"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   333
    and b: "b \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   334
    and ab: "a \<otimes> b = \<one>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   335
  shows "inv a = b"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   336
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   337
  from ab b have *: "a \<noteq> \<zero>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   338
    by (cases "a = \<zero>") simp_all
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   339
  with a have "inv a \<otimes> (a \<otimes> b) = inv a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   340
    by (simp add: ab)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   341
  with a b * show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   342
    by (simp add: m_assoc [symmetric])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   343
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   344
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   345
lemma nonzero_inverse_inverse_eq: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> inv (inv a) = a"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   346
  by (rule inverse_unique) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   347
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   348
lemma inv_1 [simp]: "inv \<one> = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   349
  by (rule inverse_unique) simp_all
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   350
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   351
lemma nonzero_inverse_mult_distrib:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   352
  assumes "a \<in> carrier R" "b \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   353
    and "a \<noteq> \<zero>" "b \<noteq> \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   354
  shows "inv (a \<otimes> b) = inv b \<otimes> inv a"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   355
proof -
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   356
  from assms have "a \<otimes> (b \<otimes> inv b) \<otimes> inv a = \<one>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   357
    by simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   358
  with assms have eq: "a \<otimes> b \<otimes> (inv b \<otimes> inv a) = \<one>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   359
    by (simp only: m_assoc m_closed inv_closed assms)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   360
  from assms show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   361
    using inverse_unique [OF _ _ eq] by simp
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   362
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   363
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   364
lemma nonzero_imp_inverse_nonzero:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   365
  assumes "a \<in> carrier R" and "a \<noteq> \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   366
  shows "inv a \<noteq> \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   367
proof
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   368
  assume *: "inv a = \<zero>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   369
  from assms have **: "\<one> = a \<otimes> inv a"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   370
    by simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   371
  also from assms have "\<dots> = \<zero>" by (simp add: *)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   372
  finally have "\<one> = \<zero>" .
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   373
  then show False by (simp add: eq_commute)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   374
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   375
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   376
lemma nonzero_divide_divide_eq_left:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   377
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   378
    a \<oslash> b \<oslash> c = a \<oslash> (b \<otimes> c)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   379
  by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   380
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   381
lemma nonzero_times_divide_eq:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   382
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   383
    b \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<otimes> (c \<oslash> d) = (a \<otimes> c) \<oslash> (b \<otimes> d)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   384
  by (simp add: m_div_def nonzero_inverse_mult_distrib m_ac integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   385
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   386
lemma nonzero_divide_divide_eq:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   387
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> d \<in> carrier R \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   388
    b \<noteq> \<zero> \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow> d \<noteq> \<zero> \<Longrightarrow> (a \<oslash> b) \<oslash> (c \<oslash> d) = (a \<otimes> d) \<oslash> (b \<otimes> c)"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   389
  by (simp add: m_div_def nonzero_inverse_mult_distrib
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   390
    nonzero_imp_inverse_nonzero nonzero_inverse_inverse_eq m_ac integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   391
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   392
lemma divide_1 [simp]: "x \<in> carrier R \<Longrightarrow> x \<oslash> \<one> = x"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   393
  by (simp add: m_div_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   394
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   395
lemma add_frac_eq:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   396
  assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   397
    and "y \<noteq> \<zero>" "z \<noteq> \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   398
  shows "x \<oslash> y \<oplus> w \<oslash> z = (x \<otimes> z \<oplus> w \<otimes> y) \<oslash> (y \<otimes> z)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   399
proof -
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   400
  from assms
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   401
  have "x \<oslash> y \<oplus> w \<oslash> z = x \<otimes> inv y \<otimes> (z \<otimes> inv z) \<oplus> w \<otimes> inv z \<otimes> (y \<otimes> inv y)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   402
    by (simp add: m_div_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   403
  also from assms have "\<dots> = (x \<otimes> z \<oplus> w \<otimes> y) \<oslash> (y \<otimes> z)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   404
    by (simp add: m_div_def nonzero_inverse_mult_distrib r_distr m_ac integral_iff del: r_inv)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   405
  finally show ?thesis .
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   406
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   407
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   408
lemma div_closed [simp]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   409
  "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> y \<noteq> \<zero> \<Longrightarrow> x \<oslash> y \<in> carrier R"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   410
  by (simp add: m_div_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   411
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   412
lemma minus_divide_left [simp]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   413
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow> \<ominus> (a \<oslash> b) = \<ominus> a \<oslash> b"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   414
  by (simp add: m_div_def l_minus)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   415
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   416
lemma diff_frac_eq:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   417
  assumes "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" "w \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   418
    and "y \<noteq> \<zero>" "z \<noteq> \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   419
  shows "x \<oslash> y \<ominus> w \<oslash> z = (x \<otimes> z \<ominus> w \<otimes> y) \<oslash> (y \<otimes> z)"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   420
  using assms by (simp add: minus_eq l_minus add_frac_eq)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   421
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   422
lemma nonzero_mult_divide_mult_cancel_left [simp]:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   423
  assumes "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   424
    and "b \<noteq> \<zero>" "c \<noteq> \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   425
  shows "(c \<otimes> a) \<oslash> (c \<otimes> b) = a \<oslash> b"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   426
proof -
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   427
  from assms have "(c \<otimes> a) \<oslash> (c \<otimes> b) = c \<otimes> a \<otimes> (inv b \<otimes> inv c)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   428
    by (simp add: m_div_def nonzero_inverse_mult_distrib integral_iff)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   429
  also from assms have "\<dots> =  a \<otimes> inv b \<otimes> (inv c \<otimes> c)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   430
    by (simp add: m_ac)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   431
  also from assms have "\<dots> =  a \<otimes> inv b"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   432
    by simp
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   433
  finally show ?thesis
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   434
    using assms by (simp add: m_div_def)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   435
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   436
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   437
lemma times_divide_eq_left [simp]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   438
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   439
    (b \<oslash> c) \<otimes> a = b \<otimes> a \<oslash> c"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   440
  by (simp add: m_div_def m_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   441
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   442
lemma times_divide_eq_right [simp]:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   443
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> c \<in> carrier R \<Longrightarrow> c \<noteq> \<zero> \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   444
    a \<otimes> (b \<oslash> c) = a \<otimes> b \<oslash> c"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   445
  by (simp add: m_div_def m_ac)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   446
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   447
lemma nonzero_power_divide:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   448
  "a \<in> carrier R \<Longrightarrow> b \<in> carrier R \<Longrightarrow> b \<noteq> \<zero> \<Longrightarrow>
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   449
    (a \<oslash> b) [^] (n::nat) = a [^] n \<oslash> b [^] n"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   450
  by (induct n) (simp_all add: nonzero_divide_divide_eq_left)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   451
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   452
lemma r_diff_distr:
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   453
  "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> z \<in> carrier R \<Longrightarrow>
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   454
    z \<otimes> (x \<ominus> y) = z \<otimes> x \<ominus> z \<otimes> y"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   455
  by (simp add: minus_eq r_distr r_minus)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   456
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   457
lemma divide_zero_left [simp]: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> \<zero> \<oslash> a = \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   458
  by (simp add: m_div_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   459
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   460
lemma divide_self: "a \<in> carrier R \<Longrightarrow> a \<noteq> \<zero> \<Longrightarrow> a \<oslash> a = \<one>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   461
  by (simp add: m_div_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   462
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   463
lemma divide_eq_0_iff:
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   464
  assumes "a \<in> carrier R" "b \<in> carrier R"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   465
    and "b \<noteq> \<zero>"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   466
  shows "a \<oslash> b = \<zero> \<longleftrightarrow> a = \<zero>"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   467
proof
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   468
  assume "a = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   469
  with assms show "a \<oslash> b = \<zero>" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   470
next
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   471
  assume "a \<oslash> b = \<zero>"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   472
  with assms have "b \<otimes> (a \<oslash> b) = b \<otimes> \<zero>" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   473
  also from assms have "b \<otimes> (a \<oslash> b) = b \<otimes> a \<oslash> b" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   474
  also from assms have "b \<otimes> a = a \<otimes> b" by (simp add: m_comm)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   475
  also from assms have "a \<otimes> b \<oslash> b = a \<otimes> (b \<oslash> b)" by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   476
  also from assms have "b \<oslash> b = \<one>" by (simp add: divide_self)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   477
  finally show "a = \<zero>" using assms by simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   478
qed
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   479
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   480
end
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   481
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   482
lemma field_class: "field (cring_class_ops::'a::field ring)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   483
  apply unfold_locales
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   484
    apply (simp_all add: cring_class_ops_def)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   485
  apply (auto simp add: Units_def)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   486
  apply (rule_tac x="1 / x" in exI)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   487
  apply simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   488
  done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   489
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   490
lemma div_class: "(x::'a::field) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   491
  apply (simp add: m_div_def m_inv_def class_simps)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   492
  apply (rule impI)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   493
  apply (rule ssubst [OF the_equality, of _ "1 / y"])
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   494
    apply simp_all
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   495
  apply (drule conjunct2)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   496
  apply (drule_tac f="\<lambda>x. x / y" in arg_cong)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   497
  apply simp
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   498
  done
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   499
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   500
interpretation field_class: field "cring_class_ops::'a::field ring"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   501
  rewrites "(\<zero>\<^bsub>cring_class_ops\<^esub>::'a) = 0"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   502
    and "(\<one>\<^bsub>cring_class_ops\<^esub>::'a) = 1"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   503
    and "(x::'a) \<oplus>\<^bsub>cring_class_ops\<^esub> y = x + y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   504
    and "(x::'a) \<otimes>\<^bsub>cring_class_ops\<^esub> y = x * y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   505
    and "\<ominus>\<^bsub>cring_class_ops\<^esub> (x::'a) = - x"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   506
    and "(x::'a) \<ominus>\<^bsub>cring_class_ops\<^esub> y = x - y"
67341
df79ef3b3a41 Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents: 67123
diff changeset
   507
    and "(x::'a) [^]\<^bsub>cring_class_ops\<^esub> n = x ^ n"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   508
    and "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>cring_class_ops\<^esub> = of_nat n"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   509
    and "((\<guillemotleft>i\<guillemotright>\<^bsub>cring_class_ops\<^esub>)::'a) = of_int i"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   510
    and "(x::'a) \<oslash>\<^bsub>cring_class_ops\<^esub> y = x / y"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 66453
diff changeset
   511
    and "(Int.of_int (numeral m)::'a) = numeral m"
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   512
  by (simp_all add: field_class class_simps div_class)
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   513
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents:
diff changeset
   514
end