src/HOL/Boolean_Algebra.thy
author haftmann
Mon, 02 Aug 2021 10:01:06 +0000
changeset 74101 d804e93ae9ff
permissions -rw-r--r--
moved theory Bit_Operations into Main corpus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74101
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Boolean_Algebra.thy
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     2
    Author:     Brian Huffman
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     3
*)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     4
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     5
section \<open>Abstract boolean Algebras\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     6
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     7
theory Boolean_Algebra
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     8
  imports Lattices
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
     9
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    10
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    11
locale boolean_algebra = conj: abel_semigroup "(\<^bold>\<sqinter>)" + disj: abel_semigroup "(\<^bold>\<squnion>)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    12
  for conj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "\<^bold>\<sqinter>" 70)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    13
    and disj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "\<^bold>\<squnion>" 65) +
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    14
  fixes compl :: "'a \<Rightarrow> 'a"  ("\<^bold>- _" [81] 80)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    15
    and zero :: "'a"  ("\<^bold>0")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    16
    and one  :: "'a"  ("\<^bold>1")
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    17
  assumes conj_disj_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<squnion> z) = (x \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    18
    and disj_conj_distrib: "x \<^bold>\<squnion> (y \<^bold>\<sqinter> z) = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (x \<^bold>\<squnion> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    19
    and conj_one_right: "x \<^bold>\<sqinter> \<^bold>1 = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    20
    and disj_zero_right: "x \<^bold>\<squnion> \<^bold>0 = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    21
    and conj_cancel_right [simp]: "x \<^bold>\<sqinter> \<^bold>- x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    22
    and disj_cancel_right [simp]: "x \<^bold>\<squnion> \<^bold>- x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    23
begin
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    24
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    25
sublocale conj: semilattice_neutr "(\<^bold>\<sqinter>)" "\<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    26
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    27
  show "x \<^bold>\<sqinter> \<^bold>1 = x" for x
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    28
    by (fact conj_one_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    29
  show "x \<^bold>\<sqinter> x = x" for x
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    30
  proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    31
    have "x \<^bold>\<sqinter> x = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    32
      by (simp add: disj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    33
    also have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    34
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    35
    also have "\<dots> = x \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    36
      by (simp only: conj_disj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    37
    also have "\<dots> = x \<^bold>\<sqinter> \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    38
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    39
    also have "\<dots> = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    40
      by (simp add: conj_one_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    41
    finally show ?thesis .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    42
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    43
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    44
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    45
sublocale disj: semilattice_neutr "(\<^bold>\<squnion>)" "\<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    46
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    47
  show "x \<^bold>\<squnion> \<^bold>0 = x" for x
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    48
    by (fact disj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    49
  show "x \<^bold>\<squnion> x = x" for x
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    50
  proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    51
    have "x \<^bold>\<squnion> x = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    52
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    53
    also have "\<dots> = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    54
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    55
    also have "\<dots> = x \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    56
      by (simp only: disj_conj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    57
    also have "\<dots> = x \<^bold>\<squnion> \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    58
      by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    59
    also have "\<dots> = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    60
      by (simp add: disj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    61
    finally show ?thesis .
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    62
  qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    63
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    64
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    65
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    66
subsection \<open>Complement\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    67
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    68
lemma complement_unique:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    69
  assumes 1: "a \<^bold>\<sqinter> x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    70
  assumes 2: "a \<^bold>\<squnion> x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    71
  assumes 3: "a \<^bold>\<sqinter> y = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    72
  assumes 4: "a \<^bold>\<squnion> y = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    73
  shows "x = y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    74
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    75
  from 1 3 have "(a \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (a \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    76
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    77
  then have "(x \<^bold>\<sqinter> a) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (y \<^bold>\<sqinter> a) \<^bold>\<squnion> (y \<^bold>\<sqinter> x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    78
    by (simp add: ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    79
  then have "x \<^bold>\<sqinter> (a \<^bold>\<squnion> y) = y \<^bold>\<sqinter> (a \<^bold>\<squnion> x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    80
    by (simp add: conj_disj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    81
  with 2 4 have "x \<^bold>\<sqinter> \<^bold>1 = y \<^bold>\<sqinter> \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    82
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    83
  then show "x = y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    84
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    85
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    86
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    87
lemma compl_unique: "x \<^bold>\<sqinter> y = \<^bold>0 \<Longrightarrow> x \<^bold>\<squnion> y = \<^bold>1 \<Longrightarrow> \<^bold>- x = y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    88
  by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    89
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    90
lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    91
proof (rule compl_unique)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    92
  show "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    93
    by (simp only: conj_cancel_right conj.commute)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    94
  show "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    95
    by (simp only: disj_cancel_right disj.commute)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    96
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    97
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    98
lemma compl_eq_compl_iff [simp]: 
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
    99
  \<open>\<^bold>- x = \<^bold>- y \<longleftrightarrow> x = y\<close>  (is \<open>?P \<longleftrightarrow> ?Q\<close>)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   100
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   101
  assume \<open>?Q\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   102
  then show ?P by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   103
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   104
  assume \<open>?P\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   105
  then have \<open>\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   106
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   107
  then show ?Q
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   108
    by simp
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   109
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   110
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   111
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   112
subsection \<open>Conjunction\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   113
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   114
lemma conj_zero_right [simp]: "x \<^bold>\<sqinter> \<^bold>0 = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   115
  using conj.left_idem conj_cancel_right by fastforce
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   116
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   117
lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   118
  by (rule compl_unique [OF conj_zero_right disj_zero_right])
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   119
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   120
lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\<sqinter> x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   121
  by (subst conj.commute) (rule conj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   122
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   123
lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   124
  by (subst conj.commute) (rule conj_cancel_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   125
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   126
lemma conj_disj_distrib2: "(y \<^bold>\<squnion> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   127
  by (simp only: conj.commute conj_disj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   128
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   129
lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   130
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   131
lemma conj_assoc: "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> z = x \<^bold>\<sqinter> (y \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   132
  by (fact ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   133
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   134
lemma conj_commute: "x \<^bold>\<sqinter> y = y \<^bold>\<sqinter> x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   135
  by (fact ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   136
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   137
lemmas conj_left_commute = conj.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   138
lemmas conj_ac = conj.assoc conj.commute conj.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   139
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   140
lemma conj_one_left: "\<^bold>1 \<^bold>\<sqinter> x = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   141
  by (fact conj.left_neutral)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   142
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   143
lemma conj_left_absorb: "x \<^bold>\<sqinter> (x \<^bold>\<sqinter> y) = x \<^bold>\<sqinter> y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   144
  by (fact conj.left_idem)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   145
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   146
lemma conj_absorb: "x \<^bold>\<sqinter> x = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   147
  by (fact conj.idem)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   148
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   149
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   150
subsection \<open>Disjunction\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   151
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   152
interpretation dual: boolean_algebra "(\<^bold>\<squnion>)" "(\<^bold>\<sqinter>)" compl \<open>\<^bold>1\<close> \<open>\<^bold>0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   153
  apply standard
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   154
       apply (rule disj_conj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   155
      apply (rule conj_disj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   156
     apply simp_all
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   157
  done
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   158
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   159
lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   160
  by (fact dual.compl_one)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   161
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   162
lemma disj_one_right [simp]: "x \<^bold>\<squnion> \<^bold>1 = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   163
  by (fact dual.conj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   164
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   165
lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\<squnion> x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   166
  by (fact dual.conj_zero_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   167
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   168
lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   169
  by (rule dual.conj_cancel_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   170
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   171
lemma disj_conj_distrib2: "(y \<^bold>\<sqinter> z) \<^bold>\<squnion> x = (y \<^bold>\<squnion> x) \<^bold>\<sqinter> (z \<^bold>\<squnion> x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   172
  by (rule dual.conj_disj_distrib2)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   173
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   174
lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   175
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   176
lemma disj_assoc: "(x \<^bold>\<squnion> y) \<^bold>\<squnion> z = x \<^bold>\<squnion> (y \<^bold>\<squnion> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   177
  by (fact ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   178
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   179
lemma disj_commute: "x \<^bold>\<squnion> y = y \<^bold>\<squnion> x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   180
  by (fact ac_simps)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   181
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   182
lemmas disj_left_commute = disj.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   183
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   184
lemmas disj_ac = disj.assoc disj.commute disj.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   185
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   186
lemma disj_zero_left: "\<^bold>0 \<^bold>\<squnion> x = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   187
  by (fact disj.left_neutral)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   188
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   189
lemma disj_left_absorb: "x \<^bold>\<squnion> (x \<^bold>\<squnion> y) = x \<^bold>\<squnion> y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   190
  by (fact disj.left_idem)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   191
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   192
lemma disj_absorb: "x \<^bold>\<squnion> x = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   193
  by (fact disj.idem)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   194
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   195
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   196
subsection \<open>De Morgan's Laws\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   197
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   198
lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\<sqinter> y) = \<^bold>- x \<^bold>\<squnion> \<^bold>- y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   199
proof (rule compl_unique)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   200
  have "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   201
    by (rule conj_disj_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   202
  also have "\<dots> = (y \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<^bold>- x)) \<^bold>\<squnion> (x \<^bold>\<sqinter> (y \<^bold>\<sqinter> \<^bold>- y))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   203
    by (simp only: conj_ac)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   204
  finally show "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   205
    by (simp only: conj_cancel_right conj_zero_right disj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   206
next
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   207
  have "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = (x \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)) \<^bold>\<sqinter> (y \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   208
    by (rule disj_conj_distrib2)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   209
  also have "\<dots> = (\<^bold>- y \<^bold>\<squnion> (x \<^bold>\<squnion> \<^bold>- x)) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> (y \<^bold>\<squnion> \<^bold>- y))"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   210
    by (simp only: disj_ac)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   211
  finally show "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   212
    by (simp only: disj_cancel_right disj_one_right conj_one_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   213
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   214
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   215
lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\<squnion> y) = \<^bold>- x \<^bold>\<sqinter> \<^bold>- y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   216
  using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   217
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   218
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   219
subsection \<open>Symmetric Difference\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   220
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   221
definition xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixr "\<^bold>\<ominus>" 65)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   222
  where "x \<^bold>\<ominus> y = (x \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   223
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   224
sublocale xor: comm_monoid xor \<open>\<^bold>0\<close>
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   225
proof
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   226
  fix x y z :: 'a
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   227
  let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   228
  have "?t \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) = ?t \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z \<^bold>\<sqinter> \<^bold>- z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   229
    by (simp only: conj_cancel_right conj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   230
  then show "(x \<^bold>\<ominus> y) \<^bold>\<ominus> z = x \<^bold>\<ominus> (y \<^bold>\<ominus> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   231
    by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   232
      (simp only: conj_disj_distribs conj_ac disj_ac)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   233
  show "x \<^bold>\<ominus> y = y \<^bold>\<ominus> x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   234
    by (simp only: xor_def conj_commute disj_commute)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   235
  show "x \<^bold>\<ominus> \<^bold>0 = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   236
    by (simp add: xor_def)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   237
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   238
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   239
lemmas xor_assoc = xor.assoc
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   240
lemmas xor_commute = xor.commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   241
lemmas xor_left_commute = xor.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   242
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   243
lemmas xor_ac = xor.assoc xor.commute xor.left_commute
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   244
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   245
lemma xor_def2: "x \<^bold>\<ominus> y = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   246
  using conj.commute conj_disj_distrib2 disj.commute xor_def by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   247
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   248
lemma xor_zero_right: "x \<^bold>\<ominus> \<^bold>0 = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   249
  by (fact xor.comm_neutral)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   250
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   251
lemma xor_zero_left: "\<^bold>0 \<^bold>\<ominus> x = x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   252
  by (fact xor.left_neutral)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   253
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   254
lemma xor_one_right [simp]: "x \<^bold>\<ominus> \<^bold>1 = \<^bold>- x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   255
  by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   256
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   257
lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\<ominus> x = \<^bold>- x"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   258
  by (subst xor_commute) (rule xor_one_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   259
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   260
lemma xor_self [simp]: "x \<^bold>\<ominus> x = \<^bold>0"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   261
  by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   262
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   263
lemma xor_left_self [simp]: "x \<^bold>\<ominus> (x \<^bold>\<ominus> y) = y"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   264
  by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   265
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   266
lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\<ominus> y = \<^bold>- (x \<^bold>\<ominus> y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   267
  by (simp add: ac_simps flip: xor_one_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   268
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   269
lemma xor_compl_right [simp]: "x \<^bold>\<ominus> \<^bold>- y = \<^bold>- (x \<^bold>\<ominus> y)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   270
  using xor_commute xor_compl_left by auto
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   271
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   272
lemma xor_cancel_right: "x \<^bold>\<ominus> \<^bold>- x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   273
  by (simp only: xor_compl_right xor_self compl_zero)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   274
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   275
lemma xor_cancel_left: "\<^bold>- x \<^bold>\<ominus> x = \<^bold>1"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   276
  by (simp only: xor_compl_left xor_self compl_zero)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   277
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   278
lemma conj_xor_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   279
proof -
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   280
  have *: "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z) =
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   281
        (y \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   282
    by (simp only: conj_cancel_right conj_zero_right disj_zero_left)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   283
  then show "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   284
    by (simp (no_asm_use) only:
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   285
        xor_def de_Morgan_disj de_Morgan_conj double_compl
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   286
        conj_disj_distribs conj_ac disj_ac)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   287
qed
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   288
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   289
lemma conj_xor_distrib2: "(y \<^bold>\<ominus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<ominus> (z \<^bold>\<sqinter> x)"
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   290
  by (simp add: conj.commute conj_xor_distrib)
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   291
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   292
lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   293
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   294
end
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   295
d804e93ae9ff moved theory Bit_Operations into Main corpus
haftmann
parents:
diff changeset
   296
end