File: int_abs.v

package info (click to toggle)
coq-math-classes 9.0.0-2
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 1,120 kB
  • sloc: python: 22; makefile: 21; sh: 2
file content (129 lines) | stat: -rw-r--r-- 4,260 bytes parent folder | download | duplicates (2)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
From Coq Require Import Ring.
Require Import
  MathClasses.interfaces.naturals MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders
  MathClasses.orders.nat_int MathClasses.theory.integers MathClasses.theory.rings MathClasses.orders.rings.

Section contents.
Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt} `{Naturals N}.

Add Ring Z : (rings.stdlib_ring_theory Z).

Lemma int_abs_unique_respectful {a b : IntAbs Z N} :
  ((=) ==> (=))%signature (int_abs Z N (ia:=a)) (int_abs Z N (ia:= b)).
Proof.
  intros x y E. unfold int_abs, int_abs_sig.
  apply (injective (naturals_to_semiring N Z)).
  destruct a as [[z1 A]|[z1 A]], b as [[z2 B]|[z2 B]].
     now rewrite A, B.
    symmetry. apply naturals.negate_to_ring. now rewrite A, B, involutive.
   apply naturals.negate_to_ring. now rewrite A, B, involutive.
  apply (injective (-)). now rewrite A, B, !involutive.
Qed.

Lemma int_abs_unique (a b : IntAbs Z N) (z : Z) :
  int_abs Z N (ia:=a) z = int_abs Z N (ia:=a) z.
Proof. now apply int_abs_unique_respectful. Qed.

Context `{!IntAbs Z N}.

Global Instance int_abs_proper: Setoid_Morphism (int_abs Z N) | 0.
Proof. split; try apply _. now apply int_abs_unique_respectful. Qed.

Context `{!SemiRing_Morphism (f : N → Z)}.

Lemma int_abs_spec x :
  { 0 ≤ x ∧ f (int_abs Z N x) = x } + { x ≤ 0 ∧ f (int_abs Z N x) = -x }.
Proof.
  unfold int_abs. destruct int_abs_sig as [[n E]|[n E]].
   left. rewrite <-E. split.
    now apply to_semiring_nonneg.
   apply (naturals.to_semiring_unique_alt _ _).
  right. split.
   apply flip_nonpos_negate. rewrite <-E. now apply to_semiring_nonneg.
  rewrite <-E. now apply (naturals.to_semiring_unique_alt _ _).
Qed.

Lemma int_abs_sig_alt x :
  { n : N | f n = x } + { n : N | f n = -x }.
Proof. destruct (int_abs_spec x) as [[??]|[??]]; eauto. Qed.

Lemma int_abs_nat n :
  int_abs Z N (f n) = n.
Proof.
  apply (injective f). destruct (int_abs_spec (f n)) as [[? E]|[? E]]; intuition.
  apply naturals.negate_to_ring. now rewrite E, involutive.
Qed.

Lemma int_abs_negate_nat n :
  int_abs Z N (-f n) = n.
Proof.
  apply (injective f). destruct (int_abs_spec (-f n)) as [[? E]|[? E]].
   symmetry. now apply naturals.negate_to_ring.
  now rewrite involutive in E.
Qed.

Lemma int_abs_negate x :
  int_abs Z N (-x) = int_abs Z N x.
Proof.
  destruct (int_abs_spec x) as [[_ E]|[_ E]].
   rewrite <-E at 1. now apply int_abs_negate_nat.
  rewrite <-E at 1. now apply int_abs_nat.
Qed.

Lemma int_abs_0_alt x : int_abs Z N x = 0 ↔ x = 0.
Proof.
  split; intros E1.
   destruct (int_abs_spec x) as [[_ E2]|[_ E2]].
    now rewrite <-E2, E1, preserves_0.
   apply flip_negate_0. now rewrite <-E2, E1, preserves_0.
  rewrite E1, <-(preserves_0 (f:=f)). now apply int_abs_nat.
Qed.

Lemma int_abs_ne_0 x : int_abs Z N x ≠ 0 ↔ x ≠ 0.
Proof. destruct (int_abs_0_alt x). intuition. Defined.

Lemma int_abs_0 : int_abs Z N 0 = 0.
Proof. now apply int_abs_0_alt. Qed.

Lemma int_abs_nonneg x :
  0 ≤ x → f (int_abs Z N x) = x.
Proof.
  intros E1. destruct (int_abs_spec x); intuition.
  setoid_replace x with (0 : Z).
   now rewrite int_abs_0, preserves_0.
  now apply (antisymmetry (≤)).
Qed.

Lemma int_abs_nonpos x :
  x ≤ 0 → f (int_abs Z N x) = -x.
Proof.
  intros E. rewrite <-int_abs_negate, int_abs_nonneg; intuition.
  now apply flip_nonpos_negate.
Qed.

Lemma int_abs_1 : int_abs Z N 1 = 1.
Proof.
  apply (injective f). rewrite preserves_1.
  apply int_abs_nonneg; solve_propholds.
Qed.

Lemma int_abs_nonneg_plus x y :
  0 ≤ x → 0 ≤ y → int_abs Z N (x + y) = int_abs Z N x + int_abs Z N y.
Proof.
  intros. apply (injective f).
  rewrite preserves_plus, !int_abs_nonneg; intuition.
  now apply nonneg_plus_compat.
Qed.

Lemma int_abs_mult x y :
  int_abs Z N (x * y) = int_abs Z N x * int_abs Z N y.
Proof.
  apply (injective f). rewrite preserves_mult.
  destruct (int_abs_spec x) as [[? Ex]|[? Ex]],
     (int_abs_spec y) as [[? Ey]|[? Ey]]; rewrite Ex, Ey.
     rewrite int_abs_nonneg. easy. now apply nonneg_mult_compat.
    rewrite int_abs_nonpos. ring. now apply nonneg_nonpos_mult.
   rewrite int_abs_nonpos. ring. now apply nonpos_nonneg_mult.
  rewrite int_abs_nonneg. ring. now apply nonpos_mult.
Qed.
End contents.