src/HOL/Decision_Procs/Commutative_Ring_Complete.thy
author kuncar
Tue, 03 Apr 2012 16:26:48 +0200
changeset 47308 9caab698dbe4
parent 44779 98d597c4193d
child 53374 a14d2a854c02
permissions -rw-r--r--
new package Lifting - initial commit
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
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     3
This theory is about of the relative completeness of method comm-ring
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     4
method.  As long as the reified atomic polynomials of type 'a pol are
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     5
in normal form, the cring method is complete.
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     6
*)
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     7
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
     8
header {* Proof of the relative completeness of method comm-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
     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
17508
c84af7f39a6b tuned theory dependencies;
wenzelm
parents: 17396
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
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    14
text {* Formalization of normal form *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    15
fun isnorm :: "'a::comm_ring pol \<Rightarrow> bool"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    16
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    17
    "isnorm (Pc c) \<longleftrightarrow> True"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    18
  | "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    19
  | "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    20
  | "isnorm (Pinj 0 P) \<longleftrightarrow> False"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    21
  | "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    22
  | "isnorm (PX P 0 Q) \<longleftrightarrow> False"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    23
  | "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
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"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
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
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
    27
(* Some helpful lemmas *)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    28
lemma norm_Pinj_0_False: "isnorm (Pinj 0 P) = False"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    29
  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
    30
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    31
lemma norm_PX_0_False: "isnorm (PX (Pc 0) i Q) = False"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    32
  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
    33
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    34
lemma norm_Pinj: "isnorm (Pinj i Q) \<Longrightarrow> isnorm Q"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    35
  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
    36
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    37
lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    38
  by (cases i) (auto, cases P, auto, case_tac pol2, auto)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    39
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    40
lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    41
  by (cases i) (auto, cases P, auto, case_tac pol2, 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
    42
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    43
lemma mkPinj_cn: "y ~= 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    44
  apply (auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    45
  apply (case_tac nat, auto simp add: norm_Pinj_0_False)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    46
  apply (case_tac pol, auto)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    47
  apply (case_tac y, auto)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    48
  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
    49
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
    50
lemma norm_PXtrans: 
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    51
  assumes A: "isnorm (PX P x Q)" 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
    52
  shows "isnorm (PX P x Q2)"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    53
proof (cases P)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    54
  case (PX p1 y p2)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    55
  with assms show ?thesis 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
    56
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    57
  case Pc
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    58
  with assms show ?thesis 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
    59
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    60
  case Pinj
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    61
  with assms show ?thesis 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
    62
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
    63
 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    64
lemma norm_PXtrans2:
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    65
  assumes "isnorm (PX P x Q)" and "isnorm Q2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    66
  shows "isnorm (PX P (Suc (n+x)) Q2)"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    67
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
    68
  case (PX p1 y p2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    69
  with assms show ?thesis 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
    70
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
    71
  case Pc
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    72
  with assms show ?thesis 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
    73
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
    74
  case Pinj
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    75
  with assms show ?thesis 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
    76
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
    77
23266
50f0a4f12ed3 tuned document;
wenzelm
parents: 22742
diff changeset
    78
text {* mkPX conserves normalizedness (@{text "_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
    79
lemma mkPX_cn: 
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
    80
  assumes "x \<noteq> 0" and "isnorm P" and "isnorm Q" 
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
    81
  shows "isnorm (mkPX P x Q)"
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
proof(cases P)
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
  case (Pc c)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    84
  with assms show ?thesis 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
    85
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
    86
  case (Pinj i Q)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    87
  with assms show ?thesis 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
    88
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
    89
  case (PX P1 y P2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    90
  with assms have Y0: "y > 0" by (cases y) auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    91
  from assms PX have "isnorm P1" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
    92
    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
    93
  from assms PX Y0 show ?thesis
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    94
    by (cases x) (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], 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
    95
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
    96
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
    97
text {* add conserves normalizedness *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    98
lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<oplus> Q)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
    99
proof (induct P Q rule: add.induct)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   100
  case (2 c i P2)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   101
  thus ?case by (cases P2) (simp_all, cases i, simp_all)
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
   102
next
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   103
  case (3 i P2 c)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   104
  thus ?case by (cases P2) (simp_all, cases i, simp_all)
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
   105
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
   106
  case (4 c P2 i Q2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   107
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   108
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   109
  with 4 show ?case
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   110
    by (cases i) (simp, cases P2, auto, case_tac pol2, 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
   111
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
   112
  case (5 P2 i Q2 c)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   113
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   114
    by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2])
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   115
  with 5 show ?case
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   116
    by (cases i) (simp, cases P2, auto, case_tac pol2, 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
   117
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
   118
  case (6 x P2 y Q2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   119
  then have Y0: "y>0" by (cases y) (auto simp add: norm_Pinj_0_False)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   120
  with 6 have X0: "x>0" by (cases x) (auto simp add: norm_Pinj_0_False)
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
   121
  have "x < y \<or> x = y \<or> x > y" by arith
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
   122
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   123
  { assume "x<y" hence "EX d. y =d + x" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   124
    then obtain d where y: "y = d + x" ..
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
   125
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   126
    note 6 X0
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
   127
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   128
    from 6 have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   129
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   130
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   131
    from 6 `x < y` y have "isnorm (Pinj d Q2)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   132
      by (cases d, simp, cases Q2, auto)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   133
    ultimately have ?case by (simp add: mkPinj_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
   134
  moreover
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
   135
  { assume "x=y"
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
   136
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   137
    from 6 have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   138
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   139
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   140
    note 6 Y0
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
   141
    moreover
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
   142
    ultimately have ?case by (simp add: mkPinj_cn) }
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
   143
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   144
  { assume "x>y" hence "EX d. x = d + y" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   145
    then obtain d where x: "x = d + y"..
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
   146
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   147
    note 6 Y0
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
   148
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   149
    from 6 have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   150
      by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   151
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   152
    from 6 `x > y` x have "isnorm (Pinj d P2)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   153
      by (cases d) (simp, cases P2, auto)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   154
    ultimately have ?case by (simp add: mkPinj_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
   155
  ultimately show ?case by blast
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
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
   157
  case (7 x P2 Q2 y R)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   158
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
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
   159
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   160
  { assume "x = 0"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   161
    with 7 have ?case by (auto simp add: norm_Pinj_0_False) }
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
   162
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   163
  { assume "x = 1"
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   164
    from 7 have "isnorm R" "isnorm P2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   165
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   166
    with 7 `x = 1` have "isnorm (R \<oplus> P2)" by simp
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   167
    with 7 `x = 1` have ?case
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   168
      by (simp add: norm_PXtrans[of Q2 y _]) }
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
   169
  moreover
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
   170
  { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   171
    then obtain d where X: "x=Suc (Suc d)" ..
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   172
    with 7 have NR: "isnorm R" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   173
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   174
    with 7 X have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   175
    with 7 X NR have "isnorm (R \<oplus> Pinj (x - 1) P2)" by simp
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   176
    with `isnorm (PX Q2 y R)` X have ?case
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   177
      by (simp add: norm_PXtrans[of Q2 y _]) }
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
   178
  ultimately show ?case by blast
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
   179
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
   180
  case (8 Q2 y R x P2)
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   181
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
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
   182
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   183
  { assume "x = 0" with 8 have ?case by (auto simp add: norm_Pinj_0_False) }
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
   184
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   185
  { assume "x = 1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   186
    with 8 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   187
    with 8 `x = 1` have "isnorm (R \<oplus> P2)" by simp
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   188
    with 8 `x = 1` have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
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
   189
  moreover
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
   190
  { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   191
    then obtain d where X: "x = Suc (Suc d)" ..
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   192
    with 8 have NR: "isnorm R" "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   193
      by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   194
    with 8 X have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   195
    with 8 `x > 1` NR have "isnorm (R \<oplus> Pinj (x - 1) P2)" by simp
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   196
    with `isnorm (PX Q2 y R)` X have ?case by (simp add: norm_PXtrans[of Q2 y _]) }
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
   197
  ultimately show ?case by blast
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
   198
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
   199
  case (9 P1 x P2 Q1 y Q2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   200
  then have Y0: "y>0" by (cases y) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   201
  with 9 have X0: "x>0" by (cases x) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   202
  with 9 have NP1: "isnorm P1" and NP2: "isnorm P2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   203
    by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2])
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   204
  with 9 have NQ1: "isnorm Q1" and NQ2: "isnorm Q2"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   205
    by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ 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
   206
  have "y < x \<or> x = y \<or> x < y" by arith
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
   207
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   208
  { assume sm1: "y < x" hence "EX d. x = d + y" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   209
    then obtain d where sm2: "x = d + y" ..
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   210
    note 9 NQ1 NP1 NP2 NQ2 sm1 sm2
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
    moreover
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
    have "isnorm (PX P1 d (Pc 0))" 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   213
    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
   214
      case (PX p1 y p2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   215
      with 9 sm1 sm2 show ?thesis by (cases d) (simp, cases p2, auto)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   216
    next
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   217
      case Pc with 9 sm1 sm2 show ?thesis by (cases d) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   218
    next
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   219
      case Pinj with 9 sm1 sm2 show ?thesis 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
   220
    qed
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   221
    ultimately have "isnorm (P2 \<oplus> Q2)" "isnorm (PX P1 (x - y) (Pc 0) \<oplus> Q1)" by auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   222
    with Y0 sm1 sm2 have ?case by (simp add: mkPX_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
   223
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   224
  { assume "x = y"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   225
    with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<oplus> Q2)" "isnorm (P1 \<oplus> Q1)" by auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   226
    with `x = y` Y0 have ?case by (simp add: mkPX_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
   227
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   228
  { assume sm1: "x < y" hence "EX d. y = d + x" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   229
    then obtain d where sm2: "y = d + x" ..
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   230
    note 9 NQ1 NP1 NP2 NQ2 sm1 sm2
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
   231
    moreover
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
   232
    have "isnorm (PX Q1 d (Pc 0))" 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   233
    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
   234
      case (PX p1 y p2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   235
      with 9 sm1 sm2 show ?thesis by (cases d) (simp, cases p2, auto)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   236
    next
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   237
      case Pc with 9 sm1 sm2 show ?thesis by (cases d) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   238
    next
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   239
      case Pinj with 9 sm1 sm2 show ?thesis 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
   240
    qed
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   241
    ultimately have "isnorm (P2 \<oplus> Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \<oplus> P1)" by auto
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   242
    with X0 sm1 sm2 have ?case by (simp add: mkPX_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
   243
  ultimately show ?case by blast
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   244
qed simp
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
   245
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   246
text {* mul concerves normalizedness *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   247
lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<otimes> Q)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   248
proof (induct P Q rule: mul.induct)
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
   249
  case (2 c i P2) thus ?case 
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   250
    by (cases P2) (simp_all, cases i, simp_all add: mkPinj_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
   251
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
   252
  case (3 i P2 c) thus ?case 
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   253
    by (cases P2) (simp_all, cases i, simp_all add: mkPinj_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
   254
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
   255
  case (4 c P2 i Q2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   256
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   257
    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
   258
  with 4 show ?case 
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   259
    by (cases "c = 0") (simp_all, cases "i = 0", simp_all add: mkPX_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
   260
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
   261
  case (5 P2 i Q2 c)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   262
  then have "isnorm P2" "isnorm Q2"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   263
    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
   264
  with 5 show ?case
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   265
    by (cases "c = 0") (simp_all, cases "i = 0", simp_all add: mkPX_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
   266
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
   267
  case (6 x P2 y Q2)
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
  have "x < y \<or> x = y \<or> x > y" by arith
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
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   270
  { assume "x < y" hence "EX d. y = d + x" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   271
    then obtain d where y: "y = d + x" ..
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
   272
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   273
    note 6
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
   274
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   275
    from 6 have "x > 0" by (cases x) (auto simp add: norm_Pinj_0_False)
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
   276
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   277
    from 6 have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   278
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   279
    from 6 `x < y` y have "isnorm (Pinj d Q2)" by (cases d) (simp, cases Q2, auto) 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   280
    ultimately have ?case by (simp add: mkPinj_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
   281
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   282
  { assume "x = y"
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
   283
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   284
    from 6 have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   285
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   286
    from 6 have "y>0" by (cases y) (auto simp add: norm_Pinj_0_False)
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
   287
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   288
    note 6
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
   289
    moreover
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
   290
    ultimately have ?case by (simp add: mkPinj_cn) }
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
   291
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   292
  { assume "x > y" hence "EX d. x = d + y" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   293
    then obtain d where x: "x = d + y" ..
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
   294
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   295
    note 6
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
   296
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   297
    from 6 have "y > 0" by (cases y) (auto simp add: norm_Pinj_0_False)
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
   298
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   299
    from 6 have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ 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
   300
    moreover
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   301
    from 6 `x > y` x have "isnorm (Pinj d P2)" by (cases d) (simp, 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
   302
    ultimately have ?case by (simp add: mkPinj_cn) }
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
   303
  ultimately show ?case by blast
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
   304
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
   305
  case (7 x P2 Q2 y R)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   306
  then have Y0: "y > 0" by (cases y) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   307
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
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
   308
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   309
  { assume "x = 0" with 7 have ?case by (auto simp add: norm_Pinj_0_False) }
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
   310
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   311
  { assume "x = 1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   312
    from 7 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   313
    with 7 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   314
    with 7 `x = 1` Y0 have ?case by (simp add: mkPX_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
   315
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   316
  { assume "x > 1" hence "EX d. x = Suc (Suc d)" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   317
    then obtain d where X: "x = Suc (Suc d)" ..
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   318
    from 7 have NR: "isnorm R" "isnorm Q2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   319
      by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
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
   320
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   321
    from 7 X have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   322
    moreover
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   323
    from 7 have "isnorm (Pinj x P2)" by (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
   324
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   325
    note 7 X
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   326
    ultimately have "isnorm (R \<otimes> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<otimes> Q2)" by auto
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   327
    with Y0 X have ?case by (simp add: mkPX_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
   328
  ultimately show ?case by blast
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
   329
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
   330
  case (8 Q2 y R x P2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   331
  then have Y0: "y>0" by (cases y) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   332
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
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
   333
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   334
  { assume "x = 0" with 8 have ?case by (auto simp add: norm_Pinj_0_False) }
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
   335
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   336
  { assume "x = 1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   337
    from 8 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   338
    with 8 `x = 1` have "isnorm (R \<otimes> P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R])
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   339
    with 8 `x = 1` Y0 have ?case by (simp add: mkPX_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
   340
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   341
  { assume "x > 1" hence "EX d. x = Suc (Suc d)" by arith
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   342
    then obtain d where X: "x = Suc (Suc d)" ..
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   343
    from 8 have NR: "isnorm R" "isnorm Q2"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   344
      by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R])
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
   345
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   346
    from 8 X have "isnorm (Pinj (x - 1) P2)" by (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
   347
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   348
    from 8 X have "isnorm (Pinj x P2)" by (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
   349
    moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   350
    note 8 X
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   351
    ultimately have "isnorm (R \<otimes> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<otimes> Q2)" by 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
   352
    with Y0 X have ?case by (simp add: mkPX_cn) }
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
  ultimately show ?case by blast
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
   354
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
   355
  case (9 P1 x P2 Q1 y Q2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   356
  from 9 have X0: "x > 0" by (cases x) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   357
  from 9 have Y0: "y > 0" by (cases y) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   358
  note 9
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
   359
  moreover
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   360
  from 9 have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x 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
   361
  moreover 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   362
  from 9 have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2])
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   363
  ultimately have "isnorm (P1 \<otimes> Q1)" "isnorm (P2 \<otimes> Q2)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   364
    "isnorm (P1 \<otimes> mkPinj 1 Q2)" "isnorm (Q1 \<otimes> 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
   365
    by (auto simp add: mkPinj_cn)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   366
  with 9 X0 Y0 have
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   367
    "isnorm (mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   368
    "isnorm (mkPX (P1 \<otimes> mkPinj (Suc 0) Q2) x (Pc 0))"  
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   369
    "isnorm (mkPX (Q1 \<otimes> 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
   370
    by (auto simp add: mkPX_cn)
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
   371
  thus ?case by (simp add: add_cn)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   372
qed simp
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
   373
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   374
text {* neg conserves normalizedness *}
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
   375
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
   376
proof (induct 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
   377
  case (Pinj i P2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   378
  then have "isnorm P2" by (simp add: norm_Pinj[of i P2])
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   379
  with Pinj show ?case 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
   380
next
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   381
  case (PX P1 x P2) note PX1 = this
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   382
  from PX have "isnorm P2" "isnorm P1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   383
    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
   384
  with PX show ?case
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   385
  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
   386
    case (PX p1 y p2)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   387
    with PX1 show ?thesis 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
   388
  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
   389
    case Pinj
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   390
    with PX1 show ?thesis by (cases x) auto
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   391
  qed (cases x, auto)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   392
qed simp
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
   393
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   394
text {* sub conserves normalizedness *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   395
lemma sub_cn: "isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<ominus> q)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   396
  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
   397
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   398
text {* sqr conserves normalizizedness *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   399
lemma sqr_cn: "isnorm P \<Longrightarrow> isnorm (sqr P)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   400
proof (induct P)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   401
  case Pc
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   402
  then show ?case by simp
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   403
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
   404
  case (Pinj i Q)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   405
  then show ?case
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   406
    by (cases Q) (auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_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
   407
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
   408
  case (PX P1 x P2)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   409
  then have "x + x ~= 0" "isnorm P2" "isnorm P1"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   410
    by (cases x, 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
   411
  with PX have "isnorm (mkPX (Pc (1 + 1) \<otimes> P1 \<otimes> mkPinj (Suc 0) P2) x (Pc 0))"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   412
      and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   413
    by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   414
  then show ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn)
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   415
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
   416
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 17508
diff changeset
   417
text {* pow conserves normalizedness *}
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   418
lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)"
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   419
proof (induct n arbitrary: P rule: less_induct)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   420
  case (less k)
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
   421
  show ?case 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   422
  proof (cases "k = 0")
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   423
    case True
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   424
    then show ?thesis by simp
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   425
  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
   426
    case False
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 33356
diff changeset
   427
    then have K2: "k div 2 < k" by (cases k) auto
44779
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   428
    from less have "isnorm (sqr P)" by (simp add: sqr_cn)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   429
    with less False K2 show ?thesis
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   430
      by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn)
98d597c4193d tuned proofs;
wenzelm
parents: 41807
diff changeset
   431
  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
   432
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
   433
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 17378
diff changeset
   434
end