├── .gitignore ├── Contracts ├── Crowdfunding.v └── Puzzle.v ├── Core ├── Automata.v └── Automata2.v ├── LICENSE ├── Libs └── bits.v ├── Makefile ├── README.md ├── _CoqProject └── options.v /.gitignore: -------------------------------------------------------------------------------- 1 | .*.aux 2 | *.a 3 | *.cma 4 | *.cmi 5 | *.cmo 6 | *.cmx 7 | *.cmxa 8 | *.cmxs 9 | *.glob 10 | *.ml.d 11 | *.ml4.d 12 | *.mli.d 13 | *.mllib.d 14 | *.mlpack.d 15 | *.native 16 | *.o 17 | *.v.d 18 | *.vio 19 | *.vo 20 | .coq-native/ 21 | .csdp.cache 22 | .lia.cache 23 | .nia.cache 24 | .nlia.cache 25 | .nra.cache 26 | csdp.cache 27 | lia.cache 28 | nia.cache 29 | nlia.cache 30 | nra.cache 31 | .coqdeps.d 32 | 33 | CoqMakefile 34 | CoqMakefile.conf 35 | -------------------------------------------------------------------------------- /Contracts/Crowdfunding.v: -------------------------------------------------------------------------------- 1 | From mathcomp 2 | Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. 3 | From fcsl 4 | Require Import pred. 5 | From scilla 6 | Require Import Automata2. 7 | From scilla 8 | Require Import options. 9 | 10 | 11 | Section Crowdfunding. 12 | (* Encoding of the Crowdfunding contract from the Scilla whitepaper *) 13 | 14 | (****************************************************** 15 | contract Crowdfunding 16 | (owner : address, 17 | max_block : uint, 18 | goal : uint) 19 | 20 | (* Mutable state description *) 21 | { 22 | backers : address => uint = []; 23 | funded : boolean = false; 24 | } 25 | 26 | (* Transition 1: Donating money *) 27 | transition Donate 28 | (sender : address, value : uint, tag : string) 29 | (* Simple filter identifying this transition *) 30 | if tag == "donate" => 31 | 32 | bs <- & backers; 33 | blk <- && block_number; 34 | let nxt_block = blk + 1 in 35 | if max_block <= nxt_block 36 | then send ( sender, amount -> 0, 37 | tag -> main, 38 | msg -> "deadline_passed">, MT) 39 | else 40 | if not (contains(bs, sender)) 41 | then let bs1 = put(sbs, ender, value) in 42 | backers := bs1; 43 | send ( sender, 44 | amount -> 0, 45 | tag -> "main", 46 | msg -> "ok">, MT) 47 | else send ( sender, 48 | amount -> 0, 49 | tag -> "main", 50 | msg -> "already_donated">, MT) 51 | 52 | (* Transition 2: Sending the funds to the owner *) 53 | transition GetFunds 54 | (sender : address, value : uint, tag : string) 55 | (* Only the owner can get the money back *) 56 | if (tag == "getfunds") && (sender == owner) => 57 | blk <- && block_number; 58 | bal <- & balance; 59 | if max_block < blk 60 | then if goal <= bal 61 | then funded := true; 62 | send ( owner, amount -> bal, 63 | tag -> "main", msg -> "funded">, MT) 64 | else send ( owner, amount -> 0, 65 | tag -> "main", msg -> "failed">, MT) 66 | else send ( owner, amount -> 0, tag -> "main", 67 | msg -> "too_early_to_claim_funds">, MT) 68 | 69 | (* Transition 3: Reclaim funds by a backer *) 70 | transition Claim 71 | (sender : address, value : uint, tag : string) 72 | if tag == "claim" => 73 | blk <- && block_number; 74 | if blk <= max_block 75 | then send ( sender, amount -> 0, tag -> "main", 76 | msg -> "too_early_to_reclaim">, MT) 77 | else bs <- & backers; 78 | bal <- & balance; 79 | if (not (contains(bs, sender))) || funded || 80 | goal <= bal 81 | then send ( sender, amount -> 0, 82 | tag -> "main", 83 | msg -> "cannot_refund">, MT) 84 | else 85 | let v = get(bs, sender) in 86 | backers := remove(bs, sender); 87 | send ( sender, amount -> v, tag -> "main", 88 | msg -> "here_is_your_money">, MT) 89 | 90 | *******************************************************) 91 | 92 | Record crowdState := CS { 93 | owner_mb_goal : address * nat * value; 94 | backers : seq (address * value); 95 | funded : bool; 96 | }. 97 | 98 | (* Administrative setters/getters *) 99 | Definition get_owner cs : address := (owner_mb_goal cs).1.1. 100 | Definition get_goal cs : value := (owner_mb_goal cs).2. 101 | Definition get_max_block cs : nat := (owner_mb_goal cs).1.2. 102 | 103 | 104 | Definition set_backers cs bs : crowdState := 105 | CS (owner_mb_goal cs) bs (funded cs). 106 | 107 | Definition set_funded cs f : crowdState := 108 | CS (owner_mb_goal cs) (backers cs) f. 109 | 110 | (* Parameters *) 111 | Variable init_owner : address. 112 | Variable init_max_block : nat. 113 | Variable init_goal : value. 114 | 115 | (* Initial state *) 116 | Definition init_state := CS (init_owner, init_max_block, init_max_block) [::] false. 117 | 118 | (*********************************************************) 119 | (********************* Transitions ***********************) 120 | (*********************************************************) 121 | 122 | (* Transition 1 *) 123 | (* 124 | transition Donate 125 | (sender : address, value : uint, tag : string) 126 | (* Simple filter identifying this transition *) 127 | if tag == "donate" => 128 | 129 | bs <- & backers; 130 | blk <- && block_number; 131 | let nxt_block = blk + 1 in 132 | if max_block <= nxt_block 133 | then send ( sender, amount -> 0, 134 | tag -> main, 135 | msg -> "deadline_passed">, MT) 136 | else 137 | if not (contains(bs, sender)) 138 | then let bs1 = put(sbs, ender, value) in 139 | backers := bs1; 140 | send ( sender, 141 | amount -> 0, 142 | tag -> "main", 143 | msg -> "ok">, MT) 144 | else send ( sender, 145 | amount -> 0, 146 | tag -> "main", 147 | msg -> "already_donated">, MT) 148 | *) 149 | 150 | (* Definition of the protocol *) 151 | Variable crowd_addr : address. 152 | 153 | Notation tft := (trans_fun_type crowdState). 154 | Definition ok_msg := [:: (0, [:: 1])]. 155 | Definition no_msg := [:: (0, [:: 0])]. 156 | 157 | Definition donate_tag := 1. 158 | Definition donate_fun : tft := fun id bal s m bc => 159 | if method m == donate_tag then 160 | let bs := backers s in 161 | let nxt_block := block_num bc + 1 in 162 | let from := sender m in 163 | if get_max_block s <= nxt_block 164 | then (s, Some (Msg 0 crowd_addr from 0 no_msg)) 165 | else if all [pred e | e.1 != from] bs 166 | (* new backer *) 167 | then let bs' := (from, val m) :: bs in 168 | let s' := set_backers s bs' in 169 | (s', Some (Msg 0 crowd_addr from 0 ok_msg)) 170 | else (s, Some (Msg 0 crowd_addr from 0 no_msg)) 171 | else (s, None). 172 | 173 | Definition donate := CTrans donate_tag donate_fun. 174 | 175 | (* Transition 2: Sending the funds to the owner *) 176 | (* 177 | transition GetFunds 178 | (sender : address, value : uint, tag : string) 179 | (* Only the owner can get the money back *) 180 | if (tag == "getfunds") && (sender == owner) => 181 | blk <- && block_number; 182 | bal <- & balance; 183 | if max_block < blk 184 | then if goal <= bal 185 | then funded := true; 186 | send ( owner, amount -> bal, 187 | tag -> "main", msg -> "funded">, MT) 188 | else send ( owner, amount -> 0, 189 | tag -> "main", msg -> "failed">, MT) 190 | else send ( owner, amount -> 0, tag -> "main", 191 | msg -> "too_early_to_claim_funds">, MT) 192 | *) 193 | 194 | Definition getfunds_tag := 2. 195 | Definition getfunds_fun : tft := fun id bal s m bc => 196 | let: from := sender m in 197 | if (method m == getfunds_tag) && (from == (get_owner s)) then 198 | let blk := block_num bc + 1 in 199 | if (get_max_block s < blk) 200 | then if get_goal s <= bal 201 | then let s' := set_funded s true in 202 | (s', Some (Msg bal crowd_addr from 0 ok_msg)) 203 | else (s, Some (Msg 0 crowd_addr from 0 no_msg)) 204 | else (s, Some (Msg 0 crowd_addr from 0 no_msg)) 205 | else (s, None). 206 | 207 | Definition get_funds := CTrans getfunds_tag getfunds_fun. 208 | 209 | (* Transition 3: Reclaim funds by a backer *) 210 | (* 211 | transition Claim 212 | (sender : address, value : uint, tag : string) 213 | if tag == "claim" => 214 | blk <- && block_number; 215 | if blk <= max_block 216 | then send ( sender, amount -> 0, tag -> "main", 217 | msg -> "too_early_to_reclaim">, MT) 218 | else bs <- & backers; 219 | bal <- & balance; 220 | if (not (contains(bs, sender))) || funded || 221 | goal <= bal 222 | then send ( sender, amount -> 0, 223 | tag -> "main", 224 | msg -> "cannot_refund">, MT) 225 | else 226 | let v = get(bs, sender) in 227 | backers := remove(bs, sender); 228 | send ( sender, amount -> v, tag -> "main", 229 | msg -> "here_is_your_money">, MT) 230 | *) 231 | 232 | Definition claim_tag := 3. 233 | Definition claim_fun : tft := fun id bal s m bc => 234 | let: from := sender m in 235 | if method m == claim_tag then 236 | let blk := block_num bc in 237 | if blk <= get_max_block s 238 | then 239 | (* Too early! *) 240 | (s, Some (Msg 0 crowd_addr from 0 no_msg)) 241 | else let bs := backers s in 242 | if [|| funded s | get_goal s <= bal] 243 | (* Cannot reimburse: campaign suceeded *) 244 | then (s, Some (Msg 0 crowd_addr from 0 no_msg)) 245 | else let n := seq.find [pred e | e.1 == from] bs in 246 | if n < size bs 247 | then let v := nth 0 (map snd bs) n in 248 | let bs' := filter [pred e | e.1 != from] bs in 249 | let s' := set_backers s bs' in 250 | (s', Some (Msg v crowd_addr from 0 ok_msg)) 251 | else 252 | (* Didn't back or already claimed *) 253 | (s, None) 254 | else (s, None). 255 | 256 | Definition claim := CTrans claim_tag claim_fun. 257 | 258 | Program Definition crowd_prot : Protocol crowdState := 259 | @CProt _ crowd_addr 0 init_state [:: donate; get_funds; claim] _. 260 | 261 | Lemma crowd_tags : tags crowd_prot = [:: 1; 2; 3]. 262 | Proof. by []. Qed. 263 | 264 | Lemma find_leq {A : eqType} (p : pred (A * nat)) (bs : seq (A * nat)) : 265 | nth 0 [seq i.2 | i <- bs] (seq.find p bs) <= sumn [seq i.2 | i <- bs]. 266 | Proof. 267 | elim: bs=>//[[a w]]bs/=Gi; case:ifP=>_/=; first by rewrite leq_addr. 268 | by rewrite (leq_trans Gi (leq_addl w _)). 269 | Qed. 270 | 271 | 272 | (***********************************************************) 273 | (** Correctness properties **) 274 | (***********************************************************) 275 | 276 | (************************************************************************ 277 | 278 | 1. The contract always has sufficient balance to reimburse everyone, 279 | unless it's successfully finished its campaign: 280 | 281 | The "funded" flag is set only if the campaign goals were reached, then 282 | all money goes to owner. Otherwise, the contract keeps all its money 283 | intact. 284 | 285 | Perhaps, we should make it stronger, adding a temporal property that 286 | one's reimbursement doesn't change. 287 | 288 | ************************************************************************) 289 | 290 | 291 | Definition balance_backed (st: cstate crowdState) : Prop := 292 | (* If the campaign not funded... *) 293 | ~~ (funded (state st)) -> 294 | (* the contract has enough funds to reimburse everyone. *) 295 | sumn (map snd (backers (state st))) <= balance st. 296 | 297 | Lemma sufficient_funds_safe : safe crowd_prot balance_backed. 298 | Proof. 299 | apply: safe_ind=>[|[id bal s]bc m M Hi]//. 300 | rewrite crowd_tags !inE in M. 301 | (* Get the exact transitions and start struggling... *) 302 | rewrite /= /apply_prot; case/orP: M; [|case/orP]=>/eqP M; rewrite M/=. 303 | 304 | (* Donate transition *) 305 | rewrite /donate_fun M eqxx. 306 | case: ifP=>/=_; [move=> {}/Hi Hi|]. 307 | - by rewrite subn0; apply: (leq_trans Hi (leq_addr (val m) bal)). 308 | case: ifP=>/=_; move=> {}/Hi Hi; last first. 309 | - by rewrite subn0; apply: (leq_trans Hi (leq_addr (val m) bal)). 310 | by rewrite subn0 /balance_backed/= in Hi *; rewrite addnC leq_add2r. 311 | 312 | (* Get funds transition. *) 313 | rewrite /getfunds_fun M eqxx. 314 | case: ifP=>//=_; case:ifP=>//=_;[|move=> {}/Hi Hi]; last first. 315 | - by rewrite subn0; apply: (leq_trans Hi (leq_addr (val m) bal)). 316 | case: ifP=>//=_; move=> {}/Hi Hi. 317 | by rewrite subn0; apply: (leq_trans Hi (leq_addr (val m) bal)). 318 | 319 | (* Claim funds back *) 320 | rewrite /claim_fun M eqxx. 321 | case: ifP=>//=_; [move=> {}/Hi Hi|]. 322 | - by rewrite subn0; apply: (leq_trans Hi (leq_addr (val m) bal)). 323 | case: ifP=>//=X. 324 | - case/orP: X; first by rewrite /balance_backed/==>->. 325 | by move=>_/Hi Z; rewrite subn0; apply: (leq_trans Z (leq_addr (val m) bal)). 326 | case: ifP=>//=G/=; move=> {}/Hi /= Hi. 327 | rewrite addnC. 328 | have H1: nth 0 [seq i.2 | i <- backers s] 329 | (seq.find [pred e | e.1 == sender m] (backers s)) <= 330 | sumn [seq i.2 | i <- backers s] by apply: find_leq. 331 | move: (leq_trans H1 Hi)=> H2. 332 | rewrite -(addnBA _ H2); clear H2. 333 | suff H3: sumn [seq i.2 | i <- backers s & [pred e | e.1 != sender m] i] <= 334 | bal - nth 0 [seq i.2 | i <- backers s] 335 | (seq.find [pred e | e.1 == sender m] (backers s)). 336 | - by apply: (leq_trans H3 (leq_addl (val m) _ )). 337 | clear M. 338 | suff H2: sumn [seq i.2 | i <- backers s & [pred e | e.1 != sender m] i] <= 339 | sumn [seq i.2 | i <- backers s] - 340 | nth 0 [seq i.2 | i <- backers s] (seq.find [pred e | e.1 == sender m] (backers s)). 341 | - by apply: (leq_trans H2); apply: leq_sub. 342 | clear Hi H1 X G bc bal id crowd_addr init_goal init_max_block init_owner. 343 | move: (backers s)=>bs{s}. 344 | elim:bs=>//[[a v]] bs/= Hi/=; case:ifP; last first. 345 | - move/negbT; case: ifP=>//= _ _; rewrite addnC -addnBA//subnn addn0. 346 | clear Hi; elim: bs=>//={a v}[[a v]]bs/=. 347 | case:ifP=>//=_ H; first by rewrite leq_add2l. 348 | by rewrite (leq_trans H (leq_addl _ _)). 349 | move/negbTE=>->/={a}; rewrite -(leq_add2l v) in Hi. 350 | by rewrite addnBA in Hi; last by apply: find_leq. 351 | Qed. 352 | 353 | (***********************************************************************) 354 | (****** Proving temporal properties ******) 355 | (***********************************************************************) 356 | 357 | (* Contribution of backer b is d is recorded in the `backers` *) 358 | Definition donated b (d : value) st := 359 | (filter [pred e | e.1 == b] (backers (state st))) == [:: (b, d)]. 360 | 361 | (* b doesn't claim its funding back *) 362 | Definition no_claims_from b (q : bstate * message) := sender q.2 != b. 363 | 364 | (************************************************************************ 365 | 366 | 2. The following lemma shows that the donation record is going to be 367 | preserved by the protocol since the moment it's been contributed, as 368 | long, as no messages from the corresponding backer b is sent to the 369 | contract. This guarantees that the contract doesn't "drop" the record 370 | about someone's donations. 371 | 372 | In conjunctions with sufficient_funds_safe (proved above) this 373 | guarantees that, if the campaign isn't funded, there is always a 374 | necessary amount on the balance to reimburse each backer, in the case 375 | of failure of the campaign. 376 | 377 | ************************************************************************) 378 | 379 | Lemma donation_preserved (b : address) (d : value): 380 | since_as_long crowd_prot (donated b d) 381 | (fun _ s' => donated b d s') (no_claims_from b). 382 | Proof. 383 | (* This is where we would need a temporal logic, but, well.. *) 384 | (* Let's prove it out of the definition. *) 385 | elim=>[|[bc m] sc Hi]st st' P R; first by rewrite /reachable'=>/=Z; subst st'. 386 | rewrite /reachable'/==>E. 387 | apply: (Hi (post (step_prot crowd_prot st bc m))); last 2 first; clear Hi. 388 | - by move=>q; move:(R q)=>{R}-R G; apply: R; apply/In_cons; right. 389 | - rewrite E; set st1 := (step_prot crowd_prot st bc m); clear E R P. 390 | by case: sc st1=>//=[[bc' m']] sc st1/=. 391 | clear E. 392 | have N: sender m != b. 393 | - suff B: no_claims_from b (bc, m) by []. 394 | by apply: R; apply/In_cons; left. 395 | case M: (method m \in tags crowd_prot); last first. 396 | - by move/negbT: M=>M; rewrite (bad_tag_step bc st M). 397 | case: st P=>id a s; rewrite /donated/==>D. 398 | case/orP: M; [|case/orP;[| rewrite orbC]]=>/eqP T; 399 | rewrite /apply_prot T/=. 400 | - rewrite /donate_fun T/=; case: ifP=>//_; case: ifP=>//_/=. 401 | by move/negbTE: N=>->. 402 | - by rewrite /getfunds_fun T/=; case: ifP=>//_; case: ifP=>//_; case:ifP. 403 | rewrite /claim_fun T/=; case:ifP=>//_; case: ifP=>//_; case: ifP=>//=X. 404 | rewrite -filter_predI/=; move/eqP:D=><-; apply/eqP. 405 | elim: (backers s)=>//=x xs Hi; rewrite Hi; clear Hi. 406 | case B: (x.1 == b)=>//=. 407 | by move/eqP: B=>?; subst b; move/negbTE: N; rewrite eq_sym=>/negbT=>->. 408 | Qed. 409 | 410 | (************************************************************************ 411 | 412 | 3. The final property: if the campaign has failed (goal hasn't been 413 | reached and the deadline has passed), every registered backer can get 414 | its donation back. 415 | 416 | TODO: formulate and prove it. 417 | 418 | ************************************************************************) 419 | 420 | Lemma can_claim_back b d st bc: 421 | (* We have donated, so the contract holds that state *) 422 | donated b d st -> 423 | (* Not funded *) 424 | ~~(funded (state st)) -> 425 | (* Balance is small: not reached the goal *) 426 | balance st < (get_goal (state st)) -> 427 | (* Block number exceeds the set number *) 428 | get_max_block (state st) < block_num bc -> 429 | (* Can emit message from b *) 430 | exists (m : message), 431 | sender m == b /\ 432 | out (step_prot crowd_prot st bc m) = Some (Msg d crowd_addr b 0 ok_msg). 433 | Proof. 434 | move=>D Nf Nb Nm. 435 | exists (Msg 0 b crowd_addr claim_tag [::]); split=>//. 436 | rewrite /step_prot. 437 | case: st D Nf Nb Nm =>id bal s/= D Nf Nb Nm. 438 | rewrite /apply_prot/=/claim_fun/=leqNgt Nm/= leqNgt Nb/=. 439 | rewrite /donated/= in D. 440 | move/negbTE: Nf=>->/=; rewrite -(has_find [pred e | e.1 == b]) has_filter. 441 | move/eqP: D=>D; rewrite D/=. 442 | congr (Some _); congr (Msg _ _ _ _ _). 443 | elim: (backers s) D=>//[[a w]]bs/=; case:ifP; first by move/eqP=>->{a}/=_; case. 444 | by move=>X Hi H; move/Hi: H=><-. 445 | Qed. 446 | 447 | 448 | (************************************************************************ 449 | 450 | 4. Can we have a logic that allows to express all these properties 451 | declaratively? Perhaps, we could do it in TLA? 452 | 453 | (This is going to be our future work.) 454 | 455 | ************************************************************************) 456 | 457 | End Crowdfunding. 458 | -------------------------------------------------------------------------------- /Contracts/Puzzle.v: -------------------------------------------------------------------------------- 1 | From mathcomp 2 | Require Import ssreflect ssrbool ssrnat eqtype seq. 3 | From scilla 4 | Require Import Automata. 5 | From scilla 6 | Require Import options. 7 | 8 | (** 9 | An encoding of the Puzzle contract from Luu et al. 10 | 11 | contract Puzzle { 12 | address public owner; 13 | bool public locked; 14 | uint public reward; 15 | bytes32 public diff; 16 | bytes public solution; 17 | 18 | function Puzzle() { //constructor 19 | owner = msg.sender; 20 | reward = msg.value; 21 | locked = false; 22 | diff = bytes32(11111); //pre-defined difficulty 23 | } 24 | 25 | function(){ //main code, runs at every invocation 26 | if (msg.sender == owner){ //update reward 27 | if (locked) throw; 28 | owner.send(reward); 29 | reward = msg.value; 30 | } else if (msg.data.length > 0){ //submit a solution 31 | if (locked) throw; 32 | if (sha256(msg.data) < diff) { 33 | msg.sender.send(reward); //send reward solution = msg.data; 34 | locked = true; 35 | } 36 | } 37 | } 38 | } 39 | 40 | *) 41 | 42 | Section Puzzle. 43 | 44 | Variable bytes32 : nat -> nat. 45 | Variable sha256 : nat -> nat. 46 | (* Variable data_tag : nat. *) 47 | 48 | Definition get_solution (pl : payload_type) := 49 | if pl is (0, n :: _) :: _ then n 50 | else 0. 51 | 52 | Record PuzzleState := 53 | PS { 54 | owner : nat; 55 | locked: bool; 56 | reward : nat; 57 | diff : nat; 58 | solution : option nat; 59 | }. 60 | 61 | Variable init_sender : nat. 62 | Variable init_reward : nat. 63 | 64 | (* Inital state after the constructor execution *) 65 | Set Warnings "-abstract-large-number". 66 | Definition s0 := PS init_sender false init_reward (bytes32 11111) None. 67 | 68 | (* Transition 1 -- changing the reward *) 69 | Definition input_filter1 : input_filter_type PuzzleState := 70 | fun m st1 => (sender m == init_sender) && (~~ locked (state st1)). 71 | 72 | Definition new_reward s r := 73 | let: PS a b _ d e := s in PS a b r d e. 74 | 75 | Definition trans_rel1 : trans_rel_type input_filter1 := 76 | fun m st1 _ s2 outs => 77 | let s1 := state st1 in 78 | s2 = new_reward s1 (value m) /\ 79 | outs = [:: (init_sender, reward s1, 0, [::])]. 80 | 81 | Definition update_reward := CTrans 1 trans_rel1. 82 | 83 | (* Transition 2 *) 84 | Definition input_filter2 : input_filter_type PuzzleState := 85 | fun m st1 => 86 | (sender m != init_sender) && (~~ locked (state st1)). 87 | 88 | Definition trans_rel2 : trans_rel_type input_filter2 := 89 | fun m st1 _ s2 outs => 90 | let s1 := state st1 in 91 | let sol := get_solution (payload m) in 92 | if sha256 sol < diff s1 93 | then s2 = PS (owner s1) true (reward s1) (diff s1) (Some sol) 94 | /\ outs = [:: (sender m, reward s1, 0, [::])] 95 | else s2 = s1 /\ outs = [::]. 96 | 97 | Definition try_solution := CTrans 2 trans_rel2. 98 | 99 | (* Puzzle protocol *) 100 | 101 | Variable puzzle_acc : nat. 102 | 103 | Program Definition puzzle_prot : Protocol PuzzleState := 104 | @CProt _ puzzle_acc init_reward s0 [:: update_reward; try_solution] _. 105 | 106 | (* Let the fun begin now *) 107 | Section PuzzleInvariants. 108 | 109 | Notation PState := (cstate PuzzleState). 110 | 111 | (* 112 | First invariant: the real balance at any time is not smaller than the reward, 113 | so the reward can be always transferred. 114 | *) 115 | 116 | 117 | Definition I_bal (st : PState) := 118 | ~~ locked (state st) -> 119 | balance st >= reward (state st). 120 | 121 | Theorem I_bal_ind : inductive puzzle_prot I_bal. 122 | Proof. 123 | split; rewrite /s0/I_bal/=; first by rewrite leqnn. 124 | move=>ts; case=>/=;[|case; last by []]=>Z [a bal s1]; 125 | subst ts=>/= Hi m pf s2 outs; last first. 126 | - rewrite /trans_rel2; case:ifP=>// _; first by case=>->. 127 | case=>->E/=/Hi=>G; subst outs. 128 | rewrite /msg_bals/=; rewrite subn0. 129 | suff X : bal <= bal + value m by apply: (leq_trans G X). 130 | by apply: leq_addr. 131 | rewrite /trans_rel1/==>[][]->E. 132 | rewrite /input_filter1/= in pf; case/andP:pf=>_ L. 133 | case: (s1) L Hi E=>x b c d e/={s1}Z/(_ Z){Z}Hi->_. 134 | by rewrite /msg_bals/msg_bal/= addn0 leq_addl. 135 | Qed. 136 | 137 | 138 | (* Since locked the balance doesn't change *) 139 | Definition is_locked st := locked (state st). 140 | 141 | Definition holds_same_balance (st st' : PState) := 142 | balance st = balance st'. 143 | 144 | (* Fact: since we're locked, the balance doesn't change 145 | (this is how it should be!) *) 146 | Theorem since_locked : since puzzle_prot is_locked holds_same_balance. 147 | Proof. 148 | move=>st1 Hl st2 [n]; elim: n st1 Hl=>/=[st _|n Hi st1 Hl [st][Hs Hr]]; first by move=><-. 149 | suff [Eb Hl']: balance st = balance st1 /\ is_locked st. 150 | - by rewrite /holds_same_balance in Hi *; rewrite -Eb{Eb}; apply: Hi=>//. 151 | clear Hr Hi st2. 152 | case: st1 Hs Hl=>acc b[ow1]lk1 rw1 d1 s1 [ts][m]/=[pf][os][s2][G]_->{st}/= Hl. 153 | rewrite /is_locked/= in Hl. rewrite Hl in pf=>{Hl}. 154 | by case: G=>/=;[|case=>//]=>Z; subst ts; simpl in pf; 155 | rewrite?/input_filter1 /input_filter2 andbC/= in pf. 156 | Qed. 157 | 158 | (* TODO: Think of more interesting temporal properties *) 159 | 160 | End PuzzleInvariants. 161 | 162 | End Puzzle. 163 | 164 | 165 | 166 | 167 | (* TODO: Implement the puzzle contract as an I/O automata *) 168 | 169 | 170 | 171 | 172 | (* TODO: Prove a bunch of atomic invariants *) 173 | 174 | 175 | -------------------------------------------------------------------------------- /Core/Automata.v: -------------------------------------------------------------------------------- 1 | From mathcomp 2 | Require Import ssreflect ssrbool ssrnat ssrfun seq. 3 | From fcsl 4 | Require Import pred. 5 | From scilla 6 | Require Import options. 7 | 8 | (* 9 | 10 | I'm not sure whether the I/O automata is the right model for this 11 | task. Perhaps, we shoudl instead rethink the model in terms of 12 | call/receive automata, which are similar to Disel's protocols in a way 13 | they handle message passing. The only difference is that here we don't 14 | have consider several communicating nodes, and instead work with just 15 | one. 16 | 17 | Thus, only "well-formed" messages get reacted to. 18 | 19 | Furthermore, any transition can end with zero to n send messages 20 | (i.e., calls), which will be sent in a non-deterministic order. 21 | 22 | This should give enough space to implement the systems of 23 | communicating contracts. 24 | 25 | Next, we can proceed on making assertions for composing contracts ans 26 | build on the research from the I/O automata community. 27 | 28 | Finally, we will have to come up with a way to represent SC automata 29 | in a language with exceptions and cost semantics, and compile it down 30 | to EVM. 31 | 32 | *) 33 | 34 | (* Implementation of a simple state *) 35 | 36 | Section State. 37 | 38 | Definition value_type := nat. 39 | Definition acc_id_type := nat. 40 | Definition method_tag_type := nat. 41 | Definition payload_type := (seq (nat * seq nat)). 42 | 43 | (* Message with a payload *) 44 | Record message := 45 | Msg { value : value_type; 46 | sender : acc_id_type; 47 | method_tag : method_tag_type; 48 | payload : payload_type 49 | }. 50 | 51 | (* State of a contract *) 52 | Record cstate (T : Type) := 53 | CState { 54 | acc_id : acc_id_type; 55 | balance : value_type; 56 | state : T 57 | }. 58 | 59 | End State. 60 | 61 | Section Protocol. 62 | 63 | (* Protocol operates on states of a predefined type T *) 64 | 65 | Section Transitions. 66 | 67 | (* State type *) 68 | Variable S: Type. 69 | 70 | Definition input_filter_type := message -> cstate S -> bool. 71 | 72 | Definition resp_type := 73 | (acc_id_type * value_type * method_tag_type * payload_type)%type. 74 | 75 | Definition msg_bal (rt : resp_type) : value_type := rt.1.1.2. 76 | Definition msg_bals outs := sumn (map msg_bal outs). 77 | 78 | Definition trans_rel_type (input_filter : input_filter_type) := 79 | forall m st1, input_filter m st1 -> 80 | forall (s2 : S) (outs : seq resp_type), Prop. 81 | 82 | (* Contract transition in the spirit of I/O automata *) 83 | Record ctransition := 84 | CTrans { 85 | (* Unique tag of a transition in the protocol *) 86 | tag : method_tag_type; 87 | 88 | (* Decidable filter on incoming messages *) 89 | input_filter : input_filter_type; 90 | 91 | (* Relation between input and output state *) 92 | trans_rel : trans_rel_type input_filter; 93 | }. 94 | 95 | End Transitions. 96 | 97 | Record Protocol (S : Type) := 98 | CProt { 99 | (*Account id *) 100 | acc_num : nat; 101 | (* Initial balance *) 102 | init_bal : nat; 103 | (* Initial state of a protocol *) 104 | init_state : S; 105 | (* Protocol comes with a set of transitions *) 106 | transitions : seq (ctransition S); 107 | (* All transitions have unique tags *) 108 | _ : uniq (map (@tag _) transitions) 109 | }. 110 | 111 | (* TODO: Should we also say something about exclusivity of 112 | transition's preconditions? *) 113 | 114 | (* TODO: Next, a program will have to be shown to refine a 115 | corresponding protocol. *) 116 | 117 | End Protocol. 118 | 119 | (***********************************************************) 120 | (*** Protocol properties ***) 121 | (***********************************************************) 122 | 123 | Section Invariants. 124 | 125 | Variable (S : Type) (p : Protocol S). 126 | 127 | Notation s0 := (init_state p). 128 | Notation acc := (acc_num p). 129 | Notation b0 := (init_bal p). 130 | Notation trans := (transitions p). 131 | 132 | (* Definition of an inductive invariant with respect to a protocol p *) 133 | (* Notice that the invariant updates the balance *) 134 | 135 | Notation State := (cstate S). 136 | 137 | Definition inductive (I : State -> Prop) := 138 | (* A. Invariant holds in the initial state *) 139 | I (CState acc b0 s0) /\ 140 | (* B. Any transition preserves the invariant *) 141 | forall ts, ts \In transitions p -> 142 | forall st1, I st1 -> 143 | forall m (pf : input_filter ts m st1) s2 outs, 144 | trans_rel pf s2 outs -> 145 | let bal' := (balance st1) - (msg_bals outs) + (value m) in 146 | I (CState acc bal' s2). 147 | 148 | (* Here we assume that we always dispatch the messages, i.e., one 149 | cannot just transfer us money. *) 150 | 151 | (* Determining whether we can reach s2 from s1 in one step *) 152 | Definition can_step (st1 st2 : State) := 153 | let: CState acc b1 s1 := st1 in 154 | exists ts m (pf : input_filter ts m st1) outs s2, 155 | [/\ ts \In transitions p, 156 | trans_rel pf s2 outs & 157 | let b2 := b1 - (msg_bals outs) + (value m) in 158 | st2 = CState acc b2 s2]. 159 | 160 | Fixpoint reachable' st1 st2 n := 161 | if n is n'.+1 162 | then exists st, can_step st1 st /\ reachable' st st2 n' 163 | else st1 = st2. 164 | 165 | Definition reachable st1 st2 := exists n, reachable' st1 st2 n. 166 | 167 | (*****************************************************) 168 | (* Some modal connectives *) 169 | (*****************************************************) 170 | 171 | (* q holds since p becomes true *) 172 | Definition since (p : State -> Prop) (q : State -> State -> Prop) := 173 | forall st, p st -> forall st', reachable st st'-> q st st'. 174 | 175 | (* TODO: can we come up with more sorts of invariants that have 176 | somewhat "temporal flavour?", e.g., in the spirit of eventual 177 | consistency. For instance, one should be able to prove that it's 178 | possible to withdraw money from th eaccount. *) 179 | 180 | End Invariants. 181 | 182 | (************************************* 183 | 184 | Great story for compositionality: verify the "library" and then link 185 | another contract against it and show that the joined interaction 186 | always produces the same traces as it were only a client contract with 187 | internal function. 188 | 189 | This is still not equivalent to the usual execution because of limited 190 | stack size. Thus, a call to library contract can fail, even if 191 | everything is implemented correctly! 192 | 193 | Wow. 194 | 195 | Just wow. 196 | 197 | **************************************) 198 | 199 | (************ 200 | 201 | Some more things to discuss or implement in the future: 202 | 203 | * Formulate other sorts invariants for outgoing messages. 204 | 205 | * Do we have to model the return value explicitly, or can we just 206 | represent it via message passing? 207 | 208 | * How to fomulate the property that shouldn't be violated by the 209 | Puzzle contract from the Luu-al:CCS'16 paper? 210 | 211 | Perhaps, for this, we need a more "Concurrent" semantics, in which 212 | message sends are disparate from the corresponding contracts. 213 | 214 | * Formalize the protocol of KoET contract and state the trace property 215 | that a previous king always gets correctly reimbursed. This would 216 | require to state properties over interactions between several 217 | parties. 218 | 219 | * Start discussing a compiler from a back-end, verified with respect 220 | to a protocol. 221 | 222 | *) 223 | 224 | -------------------------------------------------------------------------------- /Core/Automata2.v: -------------------------------------------------------------------------------- 1 | From mathcomp 2 | Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. 3 | From fcsl 4 | Require Import pred. 5 | From scilla 6 | Require Import options. 7 | 8 | (* A semantics of contracts with decidable deterministic transitions *) 9 | 10 | Section State. 11 | 12 | Definition value := nat. 13 | Definition address := nat. 14 | Definition tag := nat. 15 | Definition payload := (seq (nat * seq nat)). 16 | 17 | (* Message with a payload *) 18 | Record message := Msg { 19 | val : value; 20 | sender : address; 21 | to : address; 22 | method : tag; 23 | body : payload 24 | }. 25 | 26 | (* Augmented state of a contract *) 27 | Record cstate (T : Type) := CState { 28 | my_id : address; 29 | balance : value; 30 | state : T 31 | }. 32 | 33 | (* Global blockchain state *) 34 | Record bstate := BState { 35 | block_num : nat; 36 | (* Anything else? *) 37 | }. 38 | 39 | End State. 40 | 41 | Section Protocol. 42 | 43 | Section Transitions. 44 | 45 | (* State type *) 46 | Variable S: Type. 47 | 48 | (* A response from a transition *) 49 | Definition resp_type := (address * value * tag * payload)%type. 50 | Definition msg_bal (rt : resp_type) : value := rt.1.1.2. 51 | 52 | (* Transitions are functions *) 53 | Definition trans_fun_type := 54 | address -> value -> S -> message -> bstate -> (S * option message). 55 | 56 | (* Contract transition in the spirit of I/O automata *) 57 | Record ctransition := 58 | CTrans { 59 | (* Unique tag of a transition in the protocol *) 60 | ttag : tag; 61 | 62 | (* Funcion (bstate, state, msg) :-> (state', option msg') *) 63 | tfun : trans_fun_type; 64 | }. 65 | 66 | Definition idle_fun : trans_fun_type := fun _ _ s _ _ => (s, None). 67 | Definition idle := CTrans 0 idle_fun. 68 | 69 | End Transitions. 70 | 71 | Record Protocol (S : Type) := 72 | CProt { 73 | (*Account id *) 74 | acc : address; 75 | (* Initial balance *) 76 | init_bal : nat; 77 | (* Initial state of a protocol *) 78 | init_state : S; 79 | (* Protocol comes with a set of transitions *) 80 | transitions : seq (ctransition S); 81 | (* All transitions have unique tags *) 82 | _ : uniq (map (@ttag _) transitions) 83 | }. 84 | 85 | Definition tags {S : Type} (p : Protocol S) := 86 | map (@ttag _) (transitions p). 87 | 88 | End Protocol. 89 | 90 | Section Semantics. 91 | 92 | Variables (S : Type) (p : Protocol S). 93 | 94 | (* Blockchain schedules *) 95 | Definition schedule := seq (bstate * message). 96 | 97 | (* In a well-formed schedule block numbers only grow *) 98 | Fixpoint wf_sched (sch : schedule) := 99 | if sch is s :: sch' 100 | then let bnum := block_num s.1 in 101 | all [pred s' | bnum <= block_num s'.1] sch' && wf_sched sch' 102 | else true. 103 | 104 | Record step := 105 | Step { 106 | pre : cstate S; 107 | post : cstate S; 108 | out : option message 109 | }. 110 | 111 | Definition trace := seq step. 112 | 113 | Definition apply_prot id bal s m bc : (S * option message) := 114 | let n := seq.find (fun t => ttag t == method m) (transitions p) in 115 | let tr := nth (idle S) (transitions p) n in 116 | let f := tfun tr in 117 | f id bal s m bc. 118 | 119 | Definition step_prot pre bc m : step := 120 | let: CState id bal s := pre in 121 | let: (s', out) := apply_prot id bal s m bc in 122 | let: bal' := if out is Some m' then (bal + val m) - val m' else bal in 123 | let: post := CState id bal' s' in 124 | Step pre post out. 125 | 126 | (* Map a schedule into a trace *) 127 | Fixpoint execute (pre : cstate S) (sc: schedule) : trace := 128 | if sc is (bc, m) :: sc' 129 | then let stp := step_prot pre bc m in 130 | stp :: execute (post stp) sc' 131 | else [::]. 132 | 133 | 134 | 135 | Definition state0 := CState (acc p) (init_bal p) (init_state p). 136 | 137 | (* Execute from the initial contract state, 138 | if schedule is empty, repeat the initial state *) 139 | Definition execute0 sc := 140 | if sc is _ :: _ 141 | then execute state0 sc 142 | else [:: Step state0 state0 None]. 143 | 144 | (****************************************************************) 145 | (* Safety property: 146 | 147 | For any state s in _any_ trace, starting from the initial one, 148 | the unary property I holds (i.e., I s). 149 | 150 | *) 151 | (****************************************************************) 152 | Definition safe (I : cstate S -> Prop) : Prop := 153 | forall sc pre post out, 154 | Step pre post out \In execute0 sc -> I pre /\ I post. 155 | 156 | Lemma bad_tag_step bc m s : 157 | method m \notin (tags p) -> post (step_prot s bc m) = s. 158 | Proof. 159 | move=>M; rewrite /step_prot; case: s=>id bal s. 160 | rewrite [apply_prot _ _ _ _ _]surjective_pairing/=. 161 | rewrite /apply_prot. 162 | suff X : seq.find (fun t : ctransition S => ttag t == method m) (transitions p) = 163 | size (transitions p). 164 | by rewrite !X !nth_default ?leqnn//. 165 | move: (find_size (fun t : ctransition S => ttag t == method m) (transitions p)). 166 | rewrite leq_eqVlt; case/orP; first by move/eqP. 167 | rewrite -has_find=>G. suff X: False by []. 168 | rewrite /tags in M; clear s bal id bc. 169 | elim: (transitions p) G M=>//t ts Hi/=. 170 | case/orP; last by move=> {}/Hi Hi H; apply: Hi; rewrite inE in H; case/norP: H. 171 | by move/eqP=>->; rewrite inE eqxx. 172 | Qed. 173 | 174 | (* A strong induction principle for proving safety *) 175 | Lemma safe_ind (I : cstate S -> Prop) : 176 | I state0 -> 177 | (forall pre bc m, 178 | (method m \in tags p) -> I pre -> I (post (step_prot pre bc m))) -> 179 | safe I. 180 | Proof. 181 | move=>H1 H2; case=>[|[bc m] sc] pre pst out; first by case; case=>->->. 182 | case/In_cons=>[E|];[| rewrite -/execute]. 183 | - have Z: pst = post (step_prot state0 bc m) by rewrite -E. 184 | subst pst; split. 185 | + suff Z: pre = state0 by subst pre. 186 | move: E; rewrite /step_prot/=. 187 | by rewrite [apply_prot _ _ _ _ _]surjective_pairing; case. 188 | case M: (method m \in tags p); first by apply: H2. 189 | by rewrite bad_tag_step// M. 190 | have Hp: I (post (step_prot state0 bc m)). 191 | - case M: (method m \in tags p); first by apply: H2. 192 | by rewrite bad_tag_step// M. 193 | move: (post (step_prot state0 bc m)) Hp=>s Hi G; clear H1 bc m. 194 | elim: sc s Hi G=>//[[bc m] sc] Hi s Is. 195 | case/In_cons=>[E|];[|rewrite -/execute]; last first. 196 | - move/Hi=>H; apply: H. 197 | case M: (method m \in tags p); first by apply: H2. 198 | by rewrite bad_tag_step// M. 199 | suff Z: s = pre /\ pst = post (step_prot pre bc m). 200 | - case: Z=>??; subst pre pst; split=>//. 201 | case M: (method m \in tags p); first by apply: H2. 202 | by rewrite bad_tag_step// M. 203 | move: E; rewrite /step_prot/=; clear Is. 204 | case: s=>id bal s; rewrite [apply_prot _ _ _ _ _]surjective_pairing. 205 | by case=>->; rewrite {4}[apply_prot id bal s m bc]surjective_pairing. 206 | Qed. 207 | 208 | (*****************************************************) 209 | (* Some modal connectives *) 210 | (*****************************************************) 211 | 212 | Definition reachable' (st st' : cstate S) sc := 213 | st' = post (last (Step st st None) (execute st sc)). 214 | 215 | Definition reachable (st st' : cstate S) := 216 | exists sc, reachable' st st' sc. 217 | 218 | (* q holds since p , as long as schedule bits satisfy r. *) 219 | Definition since_as_long 220 | (p : cstate S -> Prop) 221 | (q : cstate S -> cstate S -> Prop) 222 | (r : bstate * message -> Prop) := 223 | forall sc st st', 224 | p st -> 225 | (forall b, b \In sc -> r b) -> 226 | reachable' st st' sc -> 227 | q st st'. 228 | 229 | End Semantics. 230 | 231 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Copyright (c) 2017, Ilya Sergey. 2 | All rights reserved. 3 | 4 | Redistribution and use in source and binary forms, with or without 5 | modification, are permitted provided that the following conditions are 6 | met: 7 | 8 | Redistributions of source code must retain the above copyright notice, 9 | this list of conditions and the following disclaimer. 10 | 11 | Redistributions in binary form must reproduce the above copyright 12 | notice, this list of conditions and the following disclaimer in the 13 | documentation and/or other materials provided with the distribution. 14 | 15 | THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS 16 | "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT 17 | LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR 18 | A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT 19 | HOLDER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 20 | SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT 21 | LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, 22 | DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY 23 | THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT 24 | (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE 25 | OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. -------------------------------------------------------------------------------- /Libs/bits.v: -------------------------------------------------------------------------------- 1 | From Coq 2 | Require Import Recdef. 3 | From mathcomp 4 | Require Import ssreflect ssrbool ssrnat eqtype seq tuple div. 5 | From scilla 6 | Require Import options. 7 | 8 | Definition BITS n := n.-tuple bool. 9 | Definition NIBBLE := BITS 4. 10 | Definition BYTE := BITS 8. 11 | Definition WORD := BITS 16. 12 | Definition DWORD := BITS 32. 13 | Definition QWORD := BITS 64. 14 | Definition DWORDorBYTE (d: bool) := 15 | BITS (if d then 32 else 8). 16 | 17 | 18 | Function natToBits' (n : nat) (acc : seq bool) {measure id n} : seq bool := 19 | if n is 0 then acc 20 | else let: (q, r) := (n %/ 2, n %% 2) 21 | in natToBits' q (odd r :: acc). 22 | Proof. 23 | by move=>_ _ n _; apply/ltP; apply: ltn_Pdiv=>//. 24 | Defined. 25 | 26 | Definition natToBitSeq n : seq bool := 27 | if n is 0 then [:: false] else natToBits' n [::]. 28 | 29 | Definition natToBits n : (size (natToBitSeq n)).-tuple bool := 30 | @Tuple (size (natToBitSeq n)) _ (natToBitSeq n) (eqxx _). 31 | 32 | Lemma bitSize {n1 n2} (b1 : BITS n1): 33 | n1 <= n2 -> size (nseq (n2 - n1) false ++ b1) == n2. 34 | Proof. 35 | by move=>N; rewrite size_cat size_tuple size_nseq subnK//. 36 | Qed. 37 | 38 | Function bitShift {n1 n2 : nat} (b1: BITS n1) {pf: n1 <= n2}: BITS n2 39 | := Tuple (bitSize b1 pf). 40 | 41 | (* TODO: figure out how to define it with a coercion *) 42 | Notation "# n" := (@bitShift _ _ (natToBits n) _) (at level 2). 43 | 44 | 45 | Program Definition w : WORD := #1232. 46 | 47 | (* TODO: develop the theory of 256-bit signed integers *) 48 | 49 | (* From mathcomp.algebra *) 50 | (* Require Import ssrint. *) 51 | 52 | (* Check a. *) 53 | 54 | (* Function nToB n : @tuple (size natsToBits n) bool := [tuple of natToBits n]. *) 55 | 56 | (* Check natToBits. *) 57 | 58 | (* Eval compute in natToBits' 41 [::]. *) 59 | 60 | 61 | (* Require Import Coq.Strings.String. *) 62 | 63 | 64 | (* Example a : NIBBLE := [tuple of [:: false; false; true; true]]. *) 65 | 66 | (* Fixpoint fromHex s : BITS (length s * 4) := *) 67 | (* if s is String c s *) 68 | (* then joinNibble (charToNibble c) (fromHex s) else 0. *) 69 | 70 | (* Check String. *) 71 | 72 | (* Example fortytwo := #42 : BYTE. *) 73 | 74 | (* Definition foo (t : Type) := *) 75 | (* if t is option nat then True else False. *) 76 | 77 | (* Goal option nat = nat -> False. *) 78 | (* move=> E. *) 79 | 80 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | # KNOWNTARGETS will not be passed along to CoqMakefile 2 | KNOWNTARGETS := CoqMakefile 3 | # KNOWNFILES will not get implicit targets from the final rule, and so depending on them won’t invoke the submake 4 | # Warning: These files get declared as PHONY, so any targets depending on them always get rebuilt 5 | KNOWNFILES := Makefile _CoqProject 6 | 7 | .DEFAULT_GOAL := invoke-coqmakefile 8 | 9 | CoqMakefile: Makefile _CoqProject 10 | $(COQBIN)coq_makefile -f _CoqProject -o CoqMakefile 11 | 12 | invoke-coqmakefile: CoqMakefile 13 | $(MAKE) --no-print-directory -f CoqMakefile $(filter-out $(KNOWNTARGETS),$(MAKECMDGOALS)) 14 | 15 | .PHONY: invoke-coqmakefile $(KNOWNFILES) 16 | 17 | # This should be the last rule, to handle any targets not declared above 18 | %: invoke-coqmakefile 19 | @true 20 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Smart Contract as Automata 2 | 3 | State-Transition Systems for Smart Contracts: semantics and 4 | properties. 5 | 6 | ## Building Instructions 7 | 8 | ### Requirements 9 | 10 | * [Coq v8.9](https://coq.inria.fr) 11 | * [Mathematical Components 1.6.4 - 1.7.0](https://math-comp.github.io/math-comp) (`ssreflect`) 12 | * [FCSL-PCM 1.0.0](https://github.com/imdea-software/fcsl-pcm) 13 | 14 | We recommend installing the requirements via [opam](https://opam.ocaml.org/doc/Install.html): 15 | 16 | ``` 17 | opam repo add coq-released https://coq.inria.fr/opam/released 18 | opam install coq-fcsl-pcm 19 | ``` 20 | 21 | ### Building the project 22 | 23 | ``` 24 | make clean; make -j 25 | ``` 26 | 27 | ## Project Structure 28 | 29 | * `Core/Automata2.v` - definition of the automata-based language semantics; 30 | * `Contracts/Puzzle.v` - a simple puzzle-solving game contract and its properties; 31 | * `Contracts/Crowdfunding.v` - the Crowdfunding contract and its properties; 32 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . scilla 2 | 3 | -arg -w -arg -notation-overridden 4 | 5 | options.v 6 | Libs/bits.v 7 | Core/Automata.v 8 | Core/Automata2.v 9 | Contracts/Puzzle.v 10 | Contracts/Crowdfunding.v 11 | 12 | -------------------------------------------------------------------------------- /options.v: -------------------------------------------------------------------------------- 1 | Export Set Implicit Arguments. 2 | Export Unset Strict Implicit. 3 | Export Unset Printing Implicit Defensive. 4 | Export Unset Program Cases. 5 | --------------------------------------------------------------------------------