Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Marianna Rapoport
iris-coq
Commits
678fdce7
Commit
678fdce7
authored
Feb 11, 2016
by
Ralf Jung
Browse files
switch the language over to integers
tests.v is temporarily broken
parent
6a054461
Changes
3
Hide whitespace changes
Inline
Side-by-side
heap_lang/derived.v
View file @
678fdce7
...
...
@@ -32,28 +32,28 @@ Proof.
by
rewrite
-
wp_let
//=
?gsubst_correct
?subst_empty
?to_of_val
.
Qed
.
Lemma
wp_le
E
(
n1
n2
:
nat
)
P
Q
:
Lemma
wp_le
E
(
n1
n2
:
Z
)
P
Q
:
(
n1
≤
n2
→
P
⊑
▷
Q
(
LitV
$
LitBool
true
))
→
(
n2
<
n1
→
P
⊑
▷
Q
(
LitV
$
LitBool
false
))
→
P
⊑
wp
E
(
BinOp
LeOp
(
Lit
$
Lit
Na
t
n1
)
(
Lit
$
Lit
Na
t
n2
))
Q
.
P
⊑
wp
E
(
BinOp
LeOp
(
Lit
$
Lit
In
t
n1
)
(
Lit
$
Lit
In
t
n2
))
Q
.
Proof
.
intros
.
rewrite
-
wp_bin_op
//
;
[].
destruct
(
bool_decide_reflect
(
n1
≤
n2
))
;
by
eauto
with
omega
.
Qed
.
Lemma
wp_lt
E
(
n1
n2
:
nat
)
P
Q
:
Lemma
wp_lt
E
(
n1
n2
:
Z
)
P
Q
:
(
n1
<
n2
→
P
⊑
▷
Q
(
LitV
$
LitBool
true
))
→
(
n2
≤
n1
→
P
⊑
▷
Q
(
LitV
$
LitBool
false
))
→
P
⊑
wp
E
(
BinOp
LtOp
(
Lit
$
Lit
Na
t
n1
)
(
Lit
$
Lit
Na
t
n2
))
Q
.
P
⊑
wp
E
(
BinOp
LtOp
(
Lit
$
Lit
In
t
n1
)
(
Lit
$
Lit
In
t
n2
))
Q
.
Proof
.
intros
.
rewrite
-
wp_bin_op
//
;
[].
destruct
(
bool_decide_reflect
(
n1
<
n2
))
;
by
eauto
with
omega
.
Qed
.
Lemma
wp_eq
E
(
n1
n2
:
nat
)
P
Q
:
Lemma
wp_eq
E
(
n1
n2
:
Z
)
P
Q
:
(
n1
=
n2
→
P
⊑
▷
Q
(
LitV
$
LitBool
true
))
→
(
n1
≠
n2
→
P
⊑
▷
Q
(
LitV
$
LitBool
false
))
→
P
⊑
wp
E
(
BinOp
EqOp
(
Lit
$
Lit
Na
t
n1
)
(
Lit
$
Lit
Na
t
n2
))
Q
.
P
⊑
wp
E
(
BinOp
EqOp
(
Lit
$
Lit
In
t
n1
)
(
Lit
$
Lit
In
t
n2
))
Q
.
Proof
.
intros
.
rewrite
-
wp_bin_op
//
;
[].
destruct
(
bool_decide_reflect
(
n1
=
n2
))
;
by
eauto
with
omega
.
...
...
heap_lang/heap_lang.v
View file @
678fdce7
...
...
@@ -2,13 +2,15 @@ Require Export program_logic.language prelude.strings.
Require
Import
prelude
.
gmap
.
Module
heap_lang
.
Open
Scope
Z_scope
.
(** Expressions and vals. *)
Definition
loc
:
=
positive
.
(* Really, any countable type. *)
Inductive
base_lit
:
Set
:
=
|
Lit
Na
t
(
n
:
nat
)
|
LitBool
(
b
:
bool
)
|
LitUnit
.
|
Lit
In
t
(
n
:
Z
)
|
LitBool
(
b
:
bool
)
|
LitUnit
.
Inductive
un_op
:
Set
:
=
|
NegOp
.
|
NegOp
|
MinusUnOp
.
Inductive
bin_op
:
Set
:
=
|
PlusOp
|
MinusOp
|
LeOp
|
LtOp
|
EqOp
.
...
...
@@ -152,16 +154,17 @@ Fixpoint subst (e : expr) (x : string) (v : val) : expr :=
Definition
un_op_eval
(
op
:
un_op
)
(
l
:
base_lit
)
:
option
base_lit
:
=
match
op
,
l
with
|
NegOp
,
LitBool
b
=>
Some
(
LitBool
(
negb
b
))
|
MinusUnOp
,
LitInt
n
=>
Some
(
LitInt
(-
n
))
|
_
,
_
=>
None
end
.
Definition
bin_op_eval
(
op
:
bin_op
)
(
l1
l2
:
base_lit
)
:
option
base_lit
:
=
match
op
,
l1
,
l2
with
|
PlusOp
,
Lit
Na
t
n1
,
Lit
Na
t
n2
=>
Some
$
Lit
Na
t
(
n1
+
n2
)
|
MinusOp
,
Lit
Na
t
n1
,
Lit
Na
t
n2
=>
Some
$
Lit
Na
t
(
n1
-
n2
)
|
LeOp
,
Lit
Na
t
n1
,
Lit
Na
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
≤
n2
)
|
LtOp
,
Lit
Na
t
n1
,
Lit
Na
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
<
n2
)
|
EqOp
,
Lit
Na
t
n1
,
Lit
Na
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
=
n2
)
|
PlusOp
,
Lit
In
t
n1
,
Lit
In
t
n2
=>
Some
$
Lit
In
t
(
n1
+
n2
)
|
MinusOp
,
Lit
In
t
n1
,
Lit
In
t
n2
=>
Some
$
Lit
In
t
(
n1
-
n2
)
|
LeOp
,
Lit
In
t
n1
,
Lit
In
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
≤
n2
)
|
LtOp
,
Lit
In
t
n1
,
Lit
In
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
<
n2
)
|
EqOp
,
Lit
In
t
n1
,
Lit
In
t
n2
=>
Some
$
LitBool
$
bool_decide
(
n1
=
n2
)
|
_
,
_
,
_
=>
None
end
.
...
...
heap_lang/notation.v
View file @
678fdce7
...
...
@@ -4,7 +4,7 @@ Delimit Scope lang_scope with L.
Bind
Scope
lang_scope
with
expr
val
.
Arguments
wp
{
_
_
}
_
_
%
L
_
.
Coercion
Lit
Na
t
:
nat
>->
base_lit
.
Coercion
Lit
In
t
:
Z
>->
base_lit
.
Coercion
LitBool
:
bool
>->
base_lit
.
(** No coercion from base_lit to expr. This makes is slightly easier to tell
apart language and Coq expressions. *)
...
...
@@ -22,6 +22,8 @@ Notation "e1 + e2" := (BinOp PlusOp e1%L e2%L)
(
at
level
50
,
left
associativity
)
:
lang_scope
.
Notation
"e1 - e2"
:
=
(
BinOp
MinusOp
e1
%
L
e2
%
L
)
(
at
level
50
,
left
associativity
)
:
lang_scope
.
Notation
"- e"
:
=
(
UnOp
MinusUnOp
e
%
L
)
(
at
level
35
,
right
associativity
)
:
lang_scope
.
Notation
"~ e"
:
=
(
UnOp
NegOp
e
%
L
)
(
at
level
75
,
right
associativity
)
:
lang_scope
.
Notation
"e1 ≤ e2"
:
=
(
BinOp
LeOp
e1
%
L
e2
%
L
)
(
at
level
70
)
:
lang_scope
.
Notation
"e1 < e2"
:
=
(
BinOp
LtOp
e1
%
L
e2
%
L
)
(
at
level
70
)
:
lang_scope
.
Notation
"e1 = e2"
:
=
(
BinOp
EqOp
e1
%
L
e2
%
L
)
(
at
level
70
)
:
lang_scope
.
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a 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