Section CAMPSugar.
Require Import String.
Require Import List.
Require Import EquivDec.
Require Import BasicRuntime.
Require Export CAMP.
Context {
fruntime:
foreign_runtime}.
Definition paccept :=
pconst (
drec nil).
Definition pfail :
camp :=
passert (
pconst (
dconst false)).
Definition makeSingleton (
p:
camp) :
camp :=
punop AColl p.
Definition pand (
p1 p2:
camp):=
pletEnv (
passert p1)
p2.
Definition toString p :=
punop AToString p.
Definition psome :=
pleft.
Definition pnone :=
pright.
Definition pnull := (
pconst dnone).
Definition punbrand'
p :=
punop AUnbrand p.
Definition punbrand :=
punbrand'
pit.
Definition pcast'
b p:=
pletIt (
punop (
ACast b)
p)
psome.
Definition pcast b :=
pcast'
b pit.
Definition psingleton'
p :=
pletIt (
punop ASingleton p)
psome.
Definition psingleton :=
psingleton'
pit.
Definition pWithBindings :
camp ->
camp :=
pletIt penv.
Definition pvarwith f :
camp ->
camp :=
punop (
ARec f).
Definition pvar f :
camp :=
pvarwith f pit.
Definition pdot f :
camp ->
camp :=
pletIt (
punop (
ADot f)
pit).
Definition pbdot s p :
camp := (
pletIt punbrand (
pdot s p)).
Definition pbsomedot s p :
camp := (
pletIt (
pbdot s p)
psome).
Definition varOf (
varName:
string) (
p:
camp) :=
pletEnv (
pvar varName)
p.
Definition lookup c := (
pWithBindings (
pdot c pit)).
Definition withVar (
name:
string) (
p:
camp) :
camp :=
pletIt (
lookup name)
p.
Definition withBrandedVar (
name:
string) (
p:
camp) :
camp :=
pletIt (
punbrand' (
lookup name))
p.
Definition pIS var p :=
pletIt p (
pvar var).
Example empty_binding := @
nil (
string*
data).
Definition stringConcat a b :=
pbinop ASConcat (
toString a) (
toString b).
Fixpoint pFoldRec (
camps:
list (
string*
camp)) (
init:
data) :
camp
:=
match camps with
|
nil =>
pconst init
| (
s,
p)::
ls =>
pletIt (
pletEnv
(
pdot s (
pvar "
current"))
(
pletIt (
pFoldRec ls init) (
pvar "
previous")))
p
end.
Fixpoint pMapRec (
camps:
list (
string*
camp)) :
camp
:=
match camps with
|
nil =>
pconst (
drec nil)
| (
s,
p)::
ls =>
pletEnv (
pdot s p) (
pMapRec ls)
end.
Definition pMapRecFromFold (
camps:
list (
string*
camp)) :
camp
:=
pFoldRec
(
map (
fun xy => ((
fst xy),
pdot "
current" (
snd xy)))
camps)
(
drec nil).
A reduction operator for patterns.
This is particularly useful with binary operators and sequencing constructs.
For example,
camp_reduce (pbinop AAnd) p1; p2; p3
and
camp_reduce pletIt p1; p2; p3
Fixpoint camp_reduce (
f:
camp->
camp->
camp) (
l:
list camp) :
camp
:=
match l with
|
nil =>
passert (
pconst (
dconst false))
|
p1::
l' =>
match l'
with
|
nil =>
p1
|
p2::
l'' =>
f p1 (
camp_reduce f l')
end
end.
Definition camp_binop_reduce (
b:
binOp) (
l:
list camp) :
camp :=
camp_reduce (
fun p1 p2 => (
pbinop b p1 p2))
l.
Definition camp_equiv p₁
p₂ :=
forall h c b d,
camp_eval h c p₁
b d =
camp_eval h c p₂
b d.
Theorem withAccept_id p :
camp_equiv (
pletIt pit p)
p.
Proof.
Lemma preduce_id b (
l:
list camp):
camp_equiv (
camp_reduce (
pbinop b)
l) (
camp_binop_reduce b l).
Proof.
Definition typedObject (
type:
brands) (
p:
camp) :=
pletIt (
pcast type)
p.
Definition namedObject (
varName:
string) (
type:
brands) (
p:
camp) :=
pletIt (
pcast type)
(
pletEnv (
pvar varName)
p).
Definition class (
type:
brands) (
contents:
data)
:=
dbrand type contents.
Definition returnVariables (
sl:
list string) :
camp
:=
punop (
ARecProject (
insertion_sort ODT_lt_dec sl))
penv.
Useful definitions
Definition pnow :=
pgetConstant ("
NOW").
Definition WW p :=
pletIt (
pgetConstant ("
WORLD"))
p.
Notation "
p1 ≐
p2" := (
passert (
pbinop AEq p1 p2)) (
right associativity,
at level 70,
only parsing).
Notation "‵
c" := (
pconst (
dconst c)) (
at level 0).
Definition mapall p :=
pletEnv (
punop ACount pit ≐
punop ACount (
pmap p))
(
pmap p).
Require Import ZArith.
Definition mapsnone p := (
punop ACount (
pmap p) ≐ ‵(0%
Z)).
Definition mapsone p := (
punop ACount (
pmap p) ≐ ‵(1%
Z)).
Definition notholds p :=
WW (
mapsnone p).
End CAMPSugar.
Delimit Scope camp_scope with camp.
Notation "
p1 |
p-
eq|
p2" := (
pbinop AEq p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
union|
p2" := (
pbinop AUnion p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
concat|
p2" := (
pbinop AConcat p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
mergeconcat|
p2" := (
pbinop AMergeConcat p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
and|
p2" := (
pbinop AAnd p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
or|
p2" := (
pbinop AOr p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
lt|
p2" := (
pbinop ALt p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
le|
p2" := (
pbinop ALe p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
minus|
p2" := (
pbinop AMinus p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
min|
p2" := (
pbinop AMin p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
max|
p2" := (
pbinop AMax p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
contains|
p2" := (
pbinop AContains p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 |
p-
sconcat|
p2" := (
pbinop ASConcat p1 p2) (
right associativity,
at level 70):
camp_scope.
Notation "
a '+
s+'
b" := (
stringConcat a b) (
right associativity,
at level 60) :
camp_scope.
Notation "|
p-
min-
num|(
p )" := (
punop ANumMin p) (
right associativity,
at level 70):
camp_scope.
Notation "|
p-
max-
num|(
p )" := (
punop ANumMin p) (
right associativity,
at level 70):
camp_scope.
Notation "
p1 ≐
p2" := (
passert (
pbinop AEq p1 p2)) (
right associativity,
at level 70,
only parsing):
camp_scope.
Notation "
p1 ∧
p2" := (
pand p1 p2) (
right associativity,
at level 65):
camp_scope.
Notation "…" :=
pit.
Notation "
s ↓
p" := (
pdot s p) (
right associativity,
at level 30):
camp_scope.
Notation "'
BINDINGS'" := (
pWithBindings pit) :
camp_scope.
Notation "
a '
RETURN'
b" := (
pletEnv a b) (
right associativity,
at level 30,
only parsing) :
camp_scope.
Notation "!
x" := (
punbrand'
x) (
at level 0) :
camp_scope.
Notation "
t '
COLON'
p" := (
varOf t p) (
at level 70) :
camp_scope.
Notation "
var '
IS'
p" := (
pIS var p) (
at level 71,
only parsing) :
camp_scope.
Notation "‵
c" := (
pconst (
dconst c)) (
at level 0).
Notation "#`
c" := (
pconst (
dconst c)) (
only parsing,
at level 0) :
camp_scope.
Notation "
s #->
p" := (
pdot s p) (
only parsing,
right associativity,
at level 30):
camp_scope.
Notation "
s !#->
p" := (
pbdot s p) (
only parsing,
right associativity,
at level 30):
camp_scope.
Notation "#
_" :=
pit (
only parsing):
camp_scope.
Notation "
p1 #=
p2" := (
passert (
pbinop AEq p1 p2)) (
only parsing,
right associativity,
at level 70):
camp_scope.