src/HOL/Decision_Procs/Commutative_Ring_Complete.thy
author paulson <lp15@cam.ac.uk>
Fri, 30 Dec 2022 17:48:41 +0000
changeset 76832 ab08604729a2
parent 67123 3fe40ff1b921
child 80105 2fa018321400
permissions -rw-r--r--
A further round of proof consolidation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31021
53642251a04f farewell to class recpower
haftmann
parents: 23464
diff changeset
     1
(*  Author:     Bernhard Haeupler
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
     2
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
     3
This theory is about of the relative completeness of the ring
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
     4
method.  As long as the reified atomic polynomials of type pol are
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
     5
in normal form, the ring method is complete.
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     6
*)
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     7
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     8
section \<open>Proof of the relative completeness of method comm-ring\<close>
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
     9
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    10
theory Commutative_Ring_Complete
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    11
  imports Commutative_Ring
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    12
begin
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    13
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    14
text \<open>Formalization of normal form\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
    15
fun isnorm :: "pol \<Rightarrow> bool"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    16
  where
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    17
    "isnorm (Pc c) \<longleftrightarrow> True"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    18
  | "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    19
  | "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    20
  | "isnorm (Pinj 0 P) \<longleftrightarrow> False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    21
  | "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    22
  | "isnorm (PX P 0 Q) \<longleftrightarrow> False"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    23
  | "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    24
  | "isnorm (PX (PX P1 j (Pc c)) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm (PX P1 j (Pc c)) \<and> isnorm Q"
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    25
  | "isnorm (PX P i Q) \<longleftrightarrow> isnorm P \<and> isnorm Q"
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    26
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    27
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    28
subsection \<open>Some helpful lemmas\<close>
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    29
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    30
lemma norm_Pinj_0_False: "isnorm (Pinj 0 P) = False"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    31
  by (cases P) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    32
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    33
lemma norm_PX_0_False: "isnorm (PX (Pc 0) i Q) = False"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    34
  by (cases i) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    35
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    36
lemma norm_Pinj: "isnorm (Pinj i Q) \<Longrightarrow> isnorm Q"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    37
  by (cases i) (simp add: norm_Pinj_0_False norm_PX_0_False, cases Q, auto)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    38
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    39
lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    40
  apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    41
   apply auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    42
  apply (cases P)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    43
    apply auto
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60567
diff changeset
    44
  subgoal for \<dots> pol2 by (cases pol2) auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    45
  done
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    46
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    47
lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    48
  apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    49
   apply auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    50
  apply (cases P)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    51
    apply auto
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60567
diff changeset
    52
  subgoal for \<dots> pol2 by (cases pol2) auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    53
  done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    54
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    55
lemma mkPinj_cn: "y \<noteq> 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    56
  apply (auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    57
   apply (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False)
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    58
   apply (rename_tac pol a, case_tac pol, auto)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    59
  apply (case_tac y, auto)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    60
  done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    61
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    62
lemma norm_PXtrans:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    63
  assumes A: "isnorm (PX P x Q)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    64
    and "isnorm Q2"
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    65
  shows "isnorm (PX P x Q2)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    66
proof (cases P)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    67
  case (PX p1 y p2)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    68
  with assms show ?thesis
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    69
    apply (cases x)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    70
     apply auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    71
    apply (cases p2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    72
      apply auto
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
    73
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    74
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    75
  case Pc
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
    76
  with assms show ?thesis
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
    77
    by (cases x) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    78
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    79
  case Pinj
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
    80
  with assms show ?thesis
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
    81
    by (cases x) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    82
qed
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
    83
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    84
lemma norm_PXtrans2:
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    85
  assumes "isnorm (PX P x Q)"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    86
    and "isnorm Q2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    87
  shows "isnorm (PX P (Suc (n + x)) Q2)"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    88
proof (cases P)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    89
  case (PX p1 y p2)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    90
  with assms show ?thesis
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    91
    apply (cases x)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    92
     apply auto
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    93
    apply (cases p2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
    94
      apply auto
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    95
    done
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    96
next
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    97
  case Pc
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    98
  with assms show ?thesis
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
    99
    by (cases x) auto
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   100
next
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   101
  case Pinj
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   102
  with assms show ?thesis
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   103
    by (cases x) auto
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   104
qed
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   105
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   106
text \<open>\<open>mkPX\<close> conserves normalizedness (\<open>_cn\<close>)\<close>
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   107
lemma mkPX_cn:
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   108
  assumes "x \<noteq> 0"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   109
    and "isnorm P"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   110
    and "isnorm Q"
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   111
  shows "isnorm (mkPX P x Q)"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   112
proof (cases P)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   113
  case (Pc c)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   114
  with assms show ?thesis
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   115
    by (cases x) (auto simp add: mkPinj_cn mkPX_def)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   116
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   117
  case (Pinj i Q)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   118
  with assms show ?thesis
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   119
    by (cases x) (auto simp add: mkPinj_cn mkPX_def)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   120
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   121
  case (PX P1 y P2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   122
  with assms have Y0: "y > 0"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   123
    by (cases y) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   124
  from assms PX have "isnorm P1" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   125
    by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   126
  from assms PX Y0 show ?thesis
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   127
    apply (cases x)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   128
     apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   129
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   130
      apply auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   131
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   132
qed
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   133
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   134
text \<open>\<open>add\<close> conserves normalizedness\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   135
lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>+\<rangle> Q)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   136
proof (induct P Q rule: add.induct)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   137
  case 1
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   138
  then show ?case by simp
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   139
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   140
  case (2 c i P2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   141
  then show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   142
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   143
      apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   144
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   145
     apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   146
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   147
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   148
  case (3 i P2 c)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   149
  then show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   150
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   151
      apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   152
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   153
     apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   154
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   155
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   156
  case (4 c P2 i Q2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   157
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   158
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   159
  with 4 show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   160
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   161
     apply simp
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   162
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   163
      apply auto
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60567
diff changeset
   164
    subgoal for \<dots> pol2 by (cases pol2) auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   165
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   166
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   167
  case (5 P2 i Q2 c)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   168
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   169
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   170
  with 5 show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   171
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   172
     apply simp
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   173
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   174
      apply auto
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60567
diff changeset
   175
    subgoal for \<dots> pol2 by (cases pol2) auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   176
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   177
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   178
  case (6 x P2 y Q2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   179
  then have Y0: "y > 0"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   180
    by (cases y) (auto simp add: norm_Pinj_0_False)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   181
  with 6 have X0: "x > 0"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   182
    by (cases x) (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   183
  consider "x < y" | "x = y" | "x > y" by arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   184
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   185
  proof cases
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   186
    case xy: 1
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   187
    then obtain d where y: "y = d + x"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   188
      by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   189
    from 6 have *: "isnorm P2" "isnorm Q2"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   190
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   191
    from 6 xy y have "isnorm (Pinj d Q2)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   192
      by (cases d, simp, cases Q2, auto)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   193
    with 6 X0 y * show ?thesis
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   194
      by (simp add: mkPinj_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   195
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   196
    case xy: 2
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   197
    from 6 have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   198
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   199
    with xy 6 Y0 show ?thesis
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   200
      by (simp add: mkPinj_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   201
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   202
    case xy: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   203
    then obtain d where x: "x = d + y"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   204
      by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   205
    from 6 have *: "isnorm P2" "isnorm Q2"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   206
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   207
    from 6 xy x have "isnorm (Pinj d P2)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   208
      by (cases d) (simp, cases P2, auto)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   209
    with xy 6 Y0 * x show ?thesis by (simp add: mkPinj_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   210
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   211
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   212
  case (7 x P2 Q2 y R)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   213
  consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   214
    by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   215
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   216
  proof cases
60767
ad5b4771fc19 tuned proofs;
wenzelm
parents: 60708
diff changeset
   217
    case 1
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   218
    with 7 show ?thesis
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   219
      by (auto simp add: norm_Pinj_0_False)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   220
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   221
    case x: 2
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   222
    from 7 have "isnorm R" "isnorm P2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   223
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   224
    with 7 x have "isnorm (R \<langle>+\<rangle> P2)"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   225
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   226
    with 7 x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   227
      by (simp add: norm_PXtrans[of Q2 y _])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   228
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   229
    case x: 3
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   230
    with 7 have NR: "isnorm R" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   231
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   232
    with 7 x have "isnorm (Pinj (x - 1) P2)"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   233
      by (cases P2) auto
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   234
    with 7 x NR have "isnorm (R \<langle>+\<rangle> Pinj (x - 1) P2)"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   235
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   236
    with \<open>isnorm (PX Q2 y R)\<close> x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   237
      by (simp add: norm_PXtrans[of Q2 y _])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   238
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   239
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   240
  case (8 Q2 y R x P2)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   241
  consider "x = 0" | "x = 1" | "x > 1"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   242
    by arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   243
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   244
  proof cases
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   245
    case 1
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   246
    with 8 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   247
      by (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   248
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   249
    case x: 2
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   250
    with 8 have "isnorm R" "isnorm P2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   251
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   252
    with 8 x have "isnorm (R \<langle>+\<rangle> P2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   253
      by simp
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   254
    with 8 x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   255
      by (simp add: norm_PXtrans[of Q2 y _])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   256
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   257
    case x: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   258
    then obtain d where x: "x = Suc (Suc d)" by atomize_elim arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   259
    with 8 have NR: "isnorm R" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   260
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   261
    with 8 x have "isnorm (Pinj (x - 1) P2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   262
      by (cases P2) auto
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   263
    with 8 x NR have "isnorm (R \<langle>+\<rangle> Pinj (x - 1) P2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   264
      by simp
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   265
    with \<open>isnorm (PX Q2 y R)\<close> x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   266
      by (simp add: norm_PXtrans[of Q2 y _])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   267
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   268
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   269
  case (9 P1 x P2 Q1 y Q2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   270
  then have Y0: "y > 0" by (cases y) auto
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   271
  with 9 have X0: "x > 0" by (cases x) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   272
  with 9 have NP1: "isnorm P1" and NP2: "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   273
    by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2])
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   274
  with 9 have NQ1: "isnorm Q1" and NQ2: "isnorm Q2"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   275
    by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   276
  consider "y < x" | "x = y" | "x < y" by arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   277
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   278
  proof cases
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   279
    case xy: 1
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   280
    then obtain d where x: "x = d + y"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   281
      by atomize_elim arith
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   282
    have "isnorm (PX P1 d (Pc 0))"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   283
    proof (cases P1)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   284
      case (PX p1 y p2)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   285
      with 9 xy x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   286
        by (cases d) (simp, cases p2, auto)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   287
    next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   288
      case Pc
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   289
      with 9 xy x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   290
        by (cases d) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   291
    next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   292
      case Pinj
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   293
      with 9 xy x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   294
        by (cases d) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   295
    qed
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   296
    with 9 NQ1 NP1 NP2 NQ2 xy x have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (PX P1 (x - y) (Pc 0) \<langle>+\<rangle> Q1)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   297
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   298
    with Y0 xy x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   299
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   300
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   301
    case xy: 2
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   302
    with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (P1 \<langle>+\<rangle> Q1)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   303
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   304
    with xy Y0 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   305
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   306
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   307
    case xy: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   308
    then obtain d where y: "y = d + x"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   309
      by atomize_elim arith
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   310
    have "isnorm (PX Q1 d (Pc 0))"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   311
    proof (cases Q1)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   312
      case (PX p1 y p2)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   313
      with 9 xy y show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   314
        by (cases d) (simp, cases p2, auto)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   315
    next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   316
      case Pc
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   317
      with 9 xy y show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   318
        by (cases d) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   319
    next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   320
      case Pinj
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   321
      with 9 xy y show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   322
        by (cases d) auto
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   323
    qed
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   324
    with 9 NQ1 NP1 NP2 NQ2 xy y have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \<langle>+\<rangle> P1)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   325
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   326
    with X0 xy y show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   327
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   328
  qed
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   329
qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   330
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   331
text \<open>\<open>mul\<close> concerves normalizedness\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   332
lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>*\<rangle> Q)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   333
proof (induct P Q rule: mul.induct)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   334
  case 1
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   335
  then show ?case by simp
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   336
next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   337
  case (2 c i P2)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   338
  then show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   339
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   340
      apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   341
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   342
     apply (simp_all add: mkPinj_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   343
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   344
next
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   345
  case (3 i P2 c)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   346
  then show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   347
    apply (cases P2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   348
      apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   349
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   350
     apply (simp_all add: mkPinj_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   351
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   352
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   353
  case (4 c P2 i Q2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   354
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   355
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   356
  with 4 show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   357
    apply (cases "c = 0")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   358
     apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   359
    apply (cases "i = 0")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   360
     apply (simp_all add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   361
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   362
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   363
  case (5 P2 i Q2 c)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   364
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   365
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   366
  with 5 show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   367
    apply (cases "c = 0")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   368
     apply simp_all
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   369
    apply (cases "i = 0")
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   370
     apply (simp_all add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   371
    done
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   372
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   373
  case (6 x P2 y Q2)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   374
  consider "x < y" | "x = y" | "x > y" by arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   375
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   376
  proof cases
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   377
    case xy: 1
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   378
    then obtain d where y: "y = d + x"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   379
      by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   380
    from 6 have *: "x > 0"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   381
      by (cases x) (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   382
    from 6 have **: "isnorm P2" "isnorm Q2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   383
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   384
    from 6 xy y have "isnorm (Pinj d Q2)"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   385
      apply (cases d)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   386
       apply simp
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   387
      apply (cases Q2)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   388
        apply auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   389
      done
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   390
    with 6 * ** y show ?thesis
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   391
      by (simp add: mkPinj_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   392
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   393
    case xy: 2
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   394
    from 6 have *: "isnorm P2" "isnorm Q2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   395
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   396
    from 6 have **: "y > 0"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   397
      by (cases y) (auto simp add: norm_Pinj_0_False)
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   398
    with 6 xy * ** show ?thesis
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   399
      by (simp add: mkPinj_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   400
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   401
    case xy: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   402
    then obtain d where x: "x = d + y"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   403
      by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   404
    from 6 have *: "y > 0"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   405
      by (cases y) (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   406
    from 6 have **: "isnorm P2" "isnorm Q2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   407
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   408
    from 6 xy x have "isnorm (Pinj d P2)"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   409
      apply (cases d)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   410
      apply simp
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   411
      apply (cases P2)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   412
      apply auto
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   413
      done
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   414
    with 6 xy * ** x show ?thesis
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   415
      by (simp add: mkPinj_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   416
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   417
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   418
  case (7 x P2 Q2 y R)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   419
  then have Y0: "y > 0" by (cases y) auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   420
  consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   421
    by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   422
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   423
  proof cases
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   424
    case 1
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   425
    with 7 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   426
      by (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   427
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   428
    case x: 2
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   429
    from 7 have "isnorm R" "isnorm P2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   430
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   431
    with 7 x have "isnorm (R \<langle>*\<rangle> P2)" "isnorm Q2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   432
      by (auto simp add: norm_PX1[of Q2 y R])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   433
    with 7 x Y0 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   434
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   435
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   436
    case x: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   437
    from 7 have *: "isnorm R" "isnorm Q2"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   438
      by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   439
    from 7 x have "isnorm (Pinj (x - 1) P2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   440
      by (cases P2) auto
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   441
    with 7 x * have "isnorm (R \<langle>*\<rangle> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<langle>*\<rangle> Q2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   442
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   443
    with Y0 x show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   444
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   445
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   446
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   447
  case (8 Q2 y R x P2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   448
  then have Y0: "y > 0"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   449
    by (cases y) auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   450
  consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   451
    by atomize_elim arith
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   452
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   453
  proof cases
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   454
    case 1
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   455
    with 8 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   456
      by (auto simp add: norm_Pinj_0_False)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   457
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   458
    case x: 2
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   459
    from 8 have "isnorm R" "isnorm P2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   460
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   461
    with 8 x have "isnorm (R \<langle>*\<rangle> P2)" "isnorm Q2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   462
      by (auto simp add: norm_PX1[of Q2 y R])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   463
    with 8 x Y0 show ?thesis
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   464
      by (simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   465
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   466
    case x: 3
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   467
    from 8 have *: "isnorm R" "isnorm Q2"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   468
      by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   469
    from 8 x have "isnorm (Pinj (x - 1) P2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   470
      by (cases P2) auto
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   471
    with 8 x * have "isnorm (R \<langle>*\<rangle> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<langle>*\<rangle> Q2)"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   472
      by auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60535
diff changeset
   473
    with Y0 x show ?thesis
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   474
      by (simp add: mkPX_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   475
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   476
next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   477
  case (9 P1 x P2 Q1 y Q2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   478
  from 9 have X0: "x > 0" by (cases x) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   479
  from 9 have Y0: "y > 0" by (cases y) auto
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   480
  from 9 have *: "isnorm P1" "isnorm P2"
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   481
    by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   482
  from 9 have "isnorm Q1" "isnorm Q2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   483
    by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2])
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   484
  with 9 * have "isnorm (P1 \<langle>*\<rangle> Q1)" "isnorm (P2 \<langle>*\<rangle> Q2)"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   485
    "isnorm (P1 \<langle>*\<rangle> mkPinj 1 Q2)" "isnorm (Q1 \<langle>*\<rangle> mkPinj 1 P2)"
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   486
    by (auto simp add: mkPinj_cn)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   487
  with 9 X0 Y0 have
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   488
    "isnorm (mkPX (P1 \<langle>*\<rangle> Q1) (x + y) (P2 \<langle>*\<rangle> Q2))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   489
    "isnorm (mkPX (P1 \<langle>*\<rangle> mkPinj (Suc 0) Q2) x (Pc 0))"
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   490
    "isnorm (mkPX (Q1 \<langle>*\<rangle> mkPinj (Suc 0) P2) y (Pc 0))"
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   491
    by (auto simp add: mkPX_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   492
  then show ?case
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   493
    by (simp add: add_cn)
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   494
qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   495
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   496
text \<open>neg conserves normalizedness\<close>
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   497
lemma neg_cn: "isnorm P \<Longrightarrow> isnorm (neg P)"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   498
proof (induct P)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   499
  case Pc
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   500
  then show ?case by simp
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   501
next
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   502
  case (Pinj i P2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   503
  then have "isnorm P2"
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   504
    by (simp add: norm_Pinj[of i P2])
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   505
  with Pinj show ?case
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   506
    by (cases P2) (auto, cases i, auto)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   507
next
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   508
  case (PX P1 x P2) note PX1 = this
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   509
  from PX have "isnorm P2" "isnorm P1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   510
    by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   511
  with PX show ?case
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   512
  proof (cases P1)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   513
    case (PX p1 y p2)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   514
    with PX1 show ?thesis
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   515
      by (cases x) (auto, cases p2, auto)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   516
  next
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   517
    case Pinj
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   518
    with PX1 show ?thesis
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   519
      by (cases x) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   520
  qed (cases x, auto)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   521
qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   522
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   523
text \<open>sub conserves normalizedness\<close>
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   524
lemma sub_cn: "isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<langle>-\<rangle> q)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   525
  by (simp add: sub_def add_cn neg_cn)
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   526
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   527
text \<open>sqr conserves normalizizedness\<close>
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   528
lemma sqr_cn: "isnorm P \<Longrightarrow> isnorm (sqr P)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   529
proof (induct P)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   530
  case Pc
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   531
  then show ?case by simp
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   532
next
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   533
  case (Pinj i Q)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   534
  then show ?case
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   535
    apply (cases Q)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   536
      apply (auto simp add: mkPX_cn mkPinj_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   537
    apply (cases i)
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   538
     apply (auto simp add: mkPX_cn mkPinj_cn)
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   539
    done
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   540
next
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   541
  case (PX P1 x P2)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53374
diff changeset
   542
  then have "x + x \<noteq> 0" "isnorm P2" "isnorm P1"
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   543
    by (cases x) (use PX in \<open>auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]\<close>)
64962
bf41e1109db3 Added new / improved tactics for fields and rings
berghofe
parents: 61586
diff changeset
   544
  with PX have "isnorm (mkPX (Pc (1 + 1) \<langle>*\<rangle> P1 \<langle>*\<rangle> mkPinj (Suc 0) P2) x (Pc 0))"
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   545
    and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   546
    by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   547
  then show ?case
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   548
    by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   549
qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   550
67123
3fe40ff1b921 misc tuning and modernization;
wenzelm
parents: 64962
diff changeset
   551
text \<open>\<open>pow\<close> conserves normalizedness\<close>
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   552
lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   553
proof (induct n arbitrary: P rule: less_induct)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   554
  case (less k)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 55754
diff changeset
   555
  show ?case
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   556
  proof (cases "k = 0")
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   557
    case True
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   558
    then show ?thesis by simp
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   559
  next
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   560
    case False
60535
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   561
    then have K2: "k div 2 < k"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   562
      by (cases k) auto
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   563
    from less have "isnorm (sqr P)"
25a3c522cc8f tuned proofs;
wenzelm
parents: 60533
diff changeset
   564
      by (simp add: sqr_cn)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   565
    with less False K2 show ?thesis
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   566
      by (cases "even k") (auto simp add: mul_cn elim: evenE oddE)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   567
  qed
17378
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   568
qed
105519771c67 The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff changeset
   569
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
   570
end