Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in
Toggle navigation
Menu
Open sidebar
Why3
why3
Commits
d15da6d4
Commit
d15da6d4
authored
Sep 04, 2011
by
MARCHE Claude
Browse files
more lemmas on absolute value and subtraction
parent
c74f2024
Changes
3
Hide whitespace changes
Inline
Side-by-side
theories/algebra.why
View file @
d15da6d4
...
...
@@ -32,6 +32,9 @@ theory Group
clone Assoc with type t = t, function op = op
axiom Inv_def : forall x:t. op x (inv x) = unit
lemma Inv_unit : forall x y:t. op x (inv y) = unit -> x = y
end
theory CommutativeGroup
...
...
theories/int.why
View file @
d15da6d4
...
...
@@ -19,8 +19,12 @@ theory Abs
function abs (x:int) : int = if x >= 0 then x else -x
lemma Abs_le: forall x y:int. abs x <= y <-> -y <= x <= y
lemma Abs_pos: forall x:int. abs x >= 0
lemma Abs_zero: forall x:int. abs x = 0 -> x = 0
end
theory MinMax
...
...
theories/real.why
View file @
d15da6d4
...
...
@@ -11,6 +11,8 @@ theory Real
clone export algebra.OrderedField with type t = real,
function zero = zero, function one = one, predicate (<=) = (<=)
lemma sub_zero: forall x y:real. x - y = 0.0 -> x = y
end
theory RealInfix
...
...
@@ -40,6 +42,8 @@ theory Abs
lemma Abs_pos: forall x:real. abs x >= 0.0
lemma Abs_zero: forall x:real. abs x = 0.0 -> x = 0.0
end
theory MinMax
...
...
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new file
.
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment