├── Addrs.v ├── Assets.v ├── Blocks.v ├── CTreeGrafting.v ├── CTrees.v ├── CryptoHashes.v ├── CryptoSignatures.v ├── LICENSE ├── LedgerStates.v ├── MTrees.v ├── Prelude.v ├── README.md ├── Transactions.v ├── coqcompile └── lightcrypto.pdf /Addrs.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** Addrs : Representation of addresses as 160 bit sequences. **) 6 | 7 | Require Export Prelude. 8 | 9 | Fixpoint bitseq (n:nat) : Type := 10 | match n with 11 | | 0 => unit 12 | | S n' => (bool * bitseq n')%type 13 | end. 14 | 15 | Definition bitseq_eq_dec {n} (bs1 bs2:bitseq n) : { bs1 = bs2 } + { bs1 <> bs2 }. 16 | induction n as [|n IHn]. 17 | - destruct bs1, bs2. decide equality. 18 | - destruct bs1 as [b1 bs1], bs2 as [b2 bs2]. repeat decide equality. 19 | Defined. 20 | 21 | Definition addr := bitseq 160. 22 | 23 | Definition addr_eq_dec (a1 a2: addr) : { a1 = a2 } + { a1 <> a2 }. 24 | apply bitseq_eq_dec. 25 | Defined. 26 | 27 | Fixpoint strip_bitseq_false {n:nat} {X:Type} (l:list (bitseq (S n) * X)%type) : list (bitseq n * X)%type := 28 | match l with 29 | | nil => nil 30 | | cons ((false,bs),x) l' => cons (bs,x) (strip_bitseq_false l') 31 | | cons ((true,bs),x) l' => strip_bitseq_false l' 32 | end. 33 | 34 | Fixpoint strip_bitseq_true {n:nat} {X:Type} (l:list (bitseq (S n) * X)%type) : list (bitseq n * X)%type := 35 | match l with 36 | | nil => nil 37 | | cons ((true,bs),x) l' => cons (bs,x) (strip_bitseq_true l') 38 | | cons ((false,bs),x) l' => strip_bitseq_true l' 39 | end. 40 | 41 | Lemma strip_bitseq_false_iff {n} {X} (alpha:bitseq n) (x:X) (l:list (bitseq (S n) * X)%type) : 42 | In (alpha,x) (strip_bitseq_false l) <-> In ((false,alpha),x) l. 43 | induction l as [|[[[|] beta] y] l IH]. 44 | - simpl. tauto. 45 | - simpl. split. 46 | + intros H1. right. now apply IH. 47 | + intros [H1|H1]. 48 | * inversion H1. 49 | * now apply IH. 50 | - simpl. split. 51 | + intros [H1|H1]. 52 | * left. inversion H1. reflexivity. 53 | * right. now apply IH. 54 | + intros [H1|H1]. 55 | * inversion H1. now left. 56 | * right. now apply IH. 57 | Qed. 58 | 59 | Lemma strip_bitseq_true_iff {n} {X} (alpha:bitseq n) (x:X) (l:list (bitseq (S n) * X)%type) : 60 | In (alpha,x) (strip_bitseq_true l) <-> In ((true,alpha),x) l. 61 | induction l as [|[[[|] beta] y] l IH]. 62 | - simpl. tauto. 63 | - simpl. split. 64 | + intros [H1|H1]. 65 | * left. inversion H1. reflexivity. 66 | * right. now apply IH. 67 | + intros [H1|H1]. 68 | * inversion H1. now left. 69 | * right. now apply IH. 70 | - simpl. split. 71 | + intros H1. right. now apply IH. 72 | + intros [H1|H1]. 73 | * inversion H1. 74 | * now apply IH. 75 | Qed. 76 | 77 | -------------------------------------------------------------------------------- /Assets.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** Assets: Representation of assets as triples: 6 | (hashval) id 7 | (obligation) obligations to be met in order to spend the asset 8 | (preasset) the kind of asset 9 | In addition to currency units a 'publication' preasset with no value 10 | is included to simulate 'proof of existence' applications. 11 | The obligation is (al,(m,n)) where al is a list of addresses that can sign to spend, 12 | m is the number of addresses that must sign, and n is the block height 13 | at which spending is first allowed. 14 | **) 15 | 16 | Require Export CryptoSignatures. 17 | 18 | Inductive preasset : Type := 19 | | currency : nat -> preasset (*** currency u is u units of currency ***) 20 | | publication : nat -> preasset (*** data ***) 21 | . 22 | 23 | (*** (al,(m,b)) : obligation 24 | means at least m of the addresses must sign and the block 25 | height must be at least b in order to use (e.g., spend) the asset. 26 | ***) 27 | Definition obligation : Type := prod (list addr) (prod nat nat). 28 | 29 | Definition mk_obligation (al:list addr) (m b:nat) := (al,(m,b)). 30 | 31 | Definition preasset_value (u:preasset) : option nat := 32 | match u with 33 | | currency v => Some v 34 | | _ => None 35 | end. 36 | 37 | Definition asset : Type := prod hashval (prod obligation preasset). 38 | 39 | Definition assetid (a:asset) : hashval := fst a. 40 | Definition assetobl (a:asset) : obligation := fst (snd a). 41 | Definition assetpre (a:asset) : preasset := snd (snd a). 42 | 43 | Definition asset_value (a:asset) : option nat := preasset_value (assetpre a). 44 | 45 | Definition hashpreasset (a:preasset) : hashval := 46 | match a with 47 | | currency u => hashpair (hashnat 0) (hashnat u) 48 | | publication d => hashpair (hashnat 1) (hashnat d) 49 | end. 50 | 51 | Definition hashobligation (obl:obligation) : hashval := 52 | match obl with 53 | | (al,(m,b)) => 54 | hashpair (hashlist (map hashaddr al)) (hashpair (hashnat m) (hashnat b)) 55 | end. 56 | 57 | Definition hashasset (a:asset) : hashval := 58 | match a with 59 | | (h,(obl,p)) => 60 | hashpair h (hashpair (hashobligation obl) (hashpreasset p)) 61 | end. 62 | 63 | Lemma hashpreassetinj a b : 64 | hashpreasset a = hashpreasset b -> a = b. 65 | destruct a as [u|d]; destruct b as [v|d']; 66 | try (simpl; intros H1; exfalso; apply hashpairinj in H1; destruct H1 as [H1 _]; apply hashnatinj in H1; omega). 67 | - simpl. intros H1. 68 | apply hashpairinj in H1. destruct H1 as [_ H1]. 69 | apply hashnatinj in H1. 70 | congruence. 71 | - simpl. intros H1. apply hashpairinj in H1. destruct H1 as [_ H1]. 72 | apply hashnatinj in H1. 73 | congruence. 74 | Qed. 75 | 76 | Lemma hashobligationinj (obl obl':obligation) : 77 | hashobligation obl = hashobligation obl' -> obl = obl'. 78 | destruct obl as [al [m b]]. destruct obl' as [al' [m' b']]. 79 | simpl. intros H1. 80 | apply hashpairinj in H1. destruct H1 as [H2 H3]. 81 | apply hashpairinj in H3. destruct H3 as [H4 H5]. 82 | apply hashnatinj in H4. apply hashnatinj in H5. 83 | apply hashlistinj in H2. apply hashmapinj in H2. 84 | - congruence. 85 | - exact hashaddrinj. 86 | Qed. 87 | 88 | Lemma hashassetinj a b : 89 | hashasset a = hashasset b -> a = b. 90 | destruct a as [h [obl u]]. destruct b as [k [obl' v]]. 91 | simpl. intros H1. 92 | apply hashpairinj in H1. destruct H1 as [H2 H3]. 93 | apply hashpairinj in H3. destruct H3 as [H4 H5]. 94 | apply hashobligationinj in H4. 95 | apply hashpreassetinj in H5. 96 | congruence. 97 | Qed. 98 | 99 | Definition addr_assetid : Type := prod addr hashval. 100 | Definition addr_preasset : Type := prod addr (prod obligation preasset). 101 | Definition addr_asset : Type := prod addr asset. 102 | 103 | Definition hash_addr_assetid (aa : addr_assetid) : hashval := 104 | let (alpha,h) := aa in 105 | hashpair (hashaddr alpha) h. 106 | 107 | Definition hash_addr_preasset (au : addr_preasset) : hashval := 108 | match au with 109 | | (alpha,(obl,u)) => hashlist ((hashaddr alpha)::(hashobligation obl)::(hashpreasset u)::nil) 110 | end. 111 | 112 | Lemma hash_addr_assetidinj alpha a beta b : 113 | hash_addr_assetid (alpha,a) = hash_addr_assetid (beta,b) 114 | -> alpha = beta /\ a = b. 115 | unfold hash_addr_assetid. intros H1. 116 | apply hashpairinj in H1. destruct H1 as [H2 H3]. 117 | apply hashaddrinj in H2. tauto. 118 | Qed. 119 | 120 | Lemma map_hash_addr_assetidinj inpl inpl' : 121 | map hash_addr_assetid inpl = map hash_addr_assetid inpl' 122 | -> inpl = inpl'. 123 | revert inpl'. induction inpl as [|[alpha a] inpr IH]. 124 | - intros [|[beta b] inpr']. 125 | + reflexivity. 126 | + intros H1. discriminate H1. 127 | - intros [|[beta b] inpr']. 128 | + intros H1. discriminate H1. 129 | + intros H1. 130 | change ((hash_addr_assetid (alpha, a) :: map hash_addr_assetid inpr) = 131 | (hash_addr_assetid (beta, b) :: map hash_addr_assetid inpr')) in H1. 132 | f_equal. 133 | * assert (L1: alpha = beta /\ a = b). 134 | { apply hash_addr_assetidinj. congruence. } 135 | destruct L1 as [L1a L1b]. congruence. 136 | * apply IH. congruence. 137 | Qed. 138 | 139 | Lemma hash_addr_preassetinj au bv : 140 | hash_addr_preasset au = hash_addr_preasset bv 141 | -> au = bv. 142 | destruct au as [alpha [obl u]], bv as [beta [obl' v]]. 143 | intros H1. unfold hash_addr_preasset in H1. 144 | apply hashlistinj in H1. 145 | inversion H1. apply hashaddrinj in H0. subst beta. 146 | apply hashobligationinj in H2. 147 | apply hashpreassetinj in H3. 148 | congruence. 149 | Qed. 150 | 151 | Lemma map_hash_addr_preassetinj outpl outpl' : 152 | map hash_addr_preasset outpl = map hash_addr_preasset outpl' 153 | -> outpl = outpl'. 154 | revert outpl'. induction outpl as [|[alpha u] outpr IH]. 155 | - intros [|[beta v] outpr']. 156 | + reflexivity. 157 | + intros H1. discriminate H1. 158 | - intros [|[beta v] outpr']. 159 | + intros H1. discriminate H1. 160 | + intros H1. 161 | change ((hash_addr_preasset (alpha, u) :: map hash_addr_preasset outpr) = 162 | (hash_addr_preasset (beta, v) :: map hash_addr_preasset outpr')) in H1. 163 | f_equal. 164 | * apply hash_addr_preassetinj. congruence. 165 | * apply IH. congruence. 166 | Qed. 167 | 168 | Fixpoint asset_value_out (outpl:list addr_preasset) : nat := 169 | match outpl with 170 | | (alpha,(obl,currency u))::outpr => u + asset_value_out outpr 171 | | (alpha,_)::outpr => asset_value_out outpr 172 | | nil => 0 173 | end. 174 | 175 | Lemma asset_value_out_app (l r:list addr_preasset) : 176 | asset_value_out (l ++ r) = asset_value_out l + asset_value_out r. 177 | induction l as [|[beta [obl u]] l IH]. 178 | - simpl. reflexivity. 179 | - simpl. destruct u as [u|d]; simpl; omega. 180 | Qed. 181 | 182 | Definition preasset_eq_dec (a1 a2:preasset) : {a1 = a2} + {a1 <> a2}. 183 | destruct a1 as [u|d]; destruct a2 as [v|d']; try (right; discriminate). 184 | - destruct (eq_nat_dec u v) as [D1|D1]. 185 | + left. congruence. 186 | + right. intros H1. inversion H1. tauto. 187 | - destruct (eq_nat_dec d d') as [D1|D1]. 188 | + left. congruence. 189 | + right. intros H1. inversion H1. tauto. 190 | Defined. 191 | 192 | Definition obligation_eq_dec (obl1 obl2 : obligation) : { obl1 = obl2 } + { obl1 <> obl2 }. 193 | destruct obl1 as [al1 [m1 b1]]. destruct obl2 as [al2 [m2 b2]]. 194 | destruct (eq_nat_dec m1 m2) as [D1|D1]; try (right; congruence). 195 | destruct (eq_nat_dec b1 b2) as [D2|D2]; try (right; congruence). 196 | subst m2. 197 | assert (L1: {al1 = al2} + {al1 <> al2}). 198 | { clear. revert al2. induction al1 as [|alpha1 al1 IH]; intros [|alpha2 al2]. 199 | - left. reflexivity. 200 | - right. discriminate. 201 | - right. discriminate. 202 | - destruct (addr_eq_dec alpha1 alpha2) as [D1|D1]. 203 | + destruct (IH al2) as [D2|D2]. 204 | * left. congruence. 205 | * right. congruence. 206 | + right. congruence. 207 | } 208 | destruct L1 as [D3|D3]. 209 | - left. congruence. 210 | - right. intros H1. inversion H1. contradiction (D3 H0). 211 | Defined. 212 | 213 | Definition asset_eq_dec (a1 a2:asset) : {a1 = a2} + {a1 <> a2}. 214 | destruct a1 as [h1 [obl1 u1]], a2 as [h2 [obl2 u2]]. 215 | destruct (hashval_eq_dec h1 h2). 216 | - destruct (obligation_eq_dec obl1 obl2). 217 | + destruct (preasset_eq_dec u1 u2). 218 | * left. congruence. 219 | * right. congruence. 220 | + right. congruence. 221 | - right. congruence. 222 | Defined. 223 | 224 | Fixpoint new_assets (alpha:addr) (aul:list addr_preasset) (txh:hashval) (i:nat) : list asset := 225 | match aul with 226 | | nil => nil 227 | | (beta,(obl,u))::aul' => 228 | if addr_eq_dec alpha beta then 229 | (hashpair txh (hashnat i),(obl,u))::(new_assets alpha aul' txh (S i)) 230 | else 231 | new_assets alpha aul' txh (S i) 232 | end. 233 | 234 | Fixpoint remove_assets (al:list asset) (spent:list hashval) : list asset := 235 | match al with 236 | | nil => nil 237 | | a::al' => 238 | if in_dec hashval_eq_dec (assetid a) spent then 239 | remove_assets al' spent 240 | else 241 | cons a (remove_assets al' spent) 242 | end. 243 | 244 | Fixpoint get_spent (alpha:addr) (inpl:list addr_assetid) : list hashval := 245 | match inpl with 246 | | nil => nil 247 | | cons (beta,a) inpl' => 248 | if addr_eq_dec alpha beta then 249 | cons a (get_spent alpha inpl') 250 | else 251 | get_spent alpha inpl' 252 | end. 253 | 254 | Fixpoint add_vout (txh:hashval) (outpl:list addr_preasset) (i:nat) : list (addr * asset)%type := 255 | match outpl with 256 | | nil => nil 257 | | cons (alpha,(obl,u)) outpl' => cons (alpha,(hashpair txh (hashnat i),(obl,u))) (add_vout txh outpl' (S i)) 258 | end. 259 | 260 | Lemma new_assets_iff a alpha aul txh i : 261 | In a (new_assets alpha aul txh i) 262 | <-> 263 | exists j, nth_error aul j = value (alpha,(assetobl a,assetpre a)) /\ assetid a = hashpair txh (hashnat (i+j)). 264 | destruct a as [h [obl u]]. revert i. induction aul as [|[beta [obl' v]] aur IH]. 265 | - intros i. split. 266 | + simpl. tauto. 267 | + intros [[|j] [H1 _]] ; discriminate H1. 268 | - simpl. intros i. destruct (addr_eq_dec alpha beta) as [H1|H1]; split. 269 | + intros [H2|H2]. 270 | * { exists 0. simpl. split. 271 | - inversion H2. subst beta. reflexivity. 272 | - assert (L1: i + 0 = i) by omega. 273 | rewrite L1. 274 | congruence. 275 | } 276 | * { apply IH in H2. destruct H2 as [j [H3 H4]]. exists (S j). split. 277 | - exact H3. 278 | - assert (L1: i + S j = S i + j) by omega. 279 | rewrite L1. 280 | exact H4. 281 | } 282 | + intros [[|j] [H2 H3]]; simpl in *|-*. 283 | * { left. inversion H2. f_equal. 284 | assert (L1: i + 0 = i) by omega. 285 | rewrite L1 in H3. now symmetry. 286 | } 287 | * { right. apply IH. exists j. split. 288 | - assumption. 289 | - assert (L1: S i + j = i + S j) by omega. 290 | rewrite L1. assumption. 291 | } 292 | + intros H2. apply IH in H2. destruct H2 as [j [H3 H4]]. 293 | exists (S j). split. 294 | * exact H3. 295 | * assert (L1: i + S j = S i + j) by omega. 296 | rewrite L1. assumption. 297 | + intros [[|j] [H2 H3]]. 298 | * exfalso. simpl in H2. inversion H2. congruence. 299 | * { apply IH. exists j. split. 300 | - exact H2. 301 | - assert (L1: S i + j = i + S j) by omega. 302 | rewrite L1. assumption. 303 | } 304 | Qed. 305 | 306 | Lemma remove_assets_iff a al spent : 307 | In a (remove_assets al spent) 308 | <-> 309 | In a al /\ ~In (assetid a) spent. 310 | destruct a as [h [obl u]]. induction al as [|[h' [obl' u']] ar IH]. 311 | - simpl. tauto. 312 | - simpl. destruct (in_dec hashval_eq_dec h' spent) as [H1|H1]; split. 313 | + intros H2. apply IH in H2. tauto. 314 | + intros [[H2|H2] H3]. 315 | * congruence. 316 | * apply IH. tauto. 317 | + intros [H2|H2]. 318 | * { split. 319 | - tauto. 320 | - inversion H2. subst h'. assumption. 321 | } 322 | * tauto. 323 | + intros [[H2|H2] H3]. 324 | * now left. 325 | * right. tauto. 326 | Qed. 327 | 328 | Lemma get_spent_iff h alpha inpl : 329 | In h (get_spent alpha inpl) 330 | <-> 331 | In (alpha,h) inpl. 332 | induction inpl as [|[beta h'] inpr IH]. 333 | - simpl. tauto. 334 | - simpl. destruct (addr_eq_dec alpha beta) as [H1|H1]; split. 335 | + intros [H2|H2]. 336 | * left. congruence. 337 | * right. tauto. 338 | + intros [H2|H2]. 339 | * left. inversion H2. reflexivity. 340 | * right. tauto. 341 | + tauto. 342 | + intros [H2|H2]. 343 | * congruence. 344 | * tauto. 345 | Qed. 346 | 347 | Lemma no_dups_new_assets alpha aul txh i : 348 | no_dups (new_assets alpha aul txh i). 349 | revert i. induction aul as [|[beta [obl v]] aur IH]; intros i. 350 | - simpl. apply no_dups_nil. 351 | - simpl. destruct (addr_eq_dec alpha beta) as [H1|H1]. 352 | + apply no_dups_cons. 353 | * intros H2. apply new_assets_iff in H2. destruct H2 as [j [H3 H4]]. 354 | apply hashpairinj in H4. destruct H4 as [_ H4]. 355 | apply hashnatinj in H4. omega. 356 | * apply IH. 357 | + apply IH. 358 | Qed. 359 | 360 | Lemma no_dups_remove_assets al spent : 361 | no_dups al -> no_dups (remove_assets al spent). 362 | intros H. induction H as [|[h [obl u]] al H1 H2 IH]. 363 | - simpl. apply no_dups_nil. 364 | - simpl. destruct (in_dec hashval_eq_dec h spent) as [H3|H3]. 365 | + assumption. 366 | + apply no_dups_cons. 367 | * intros H4. apply remove_assets_iff in H4. tauto. 368 | * assumption. 369 | Qed. 370 | 371 | Lemma remove_assets_noint_eq (al:list asset) (rem:list hashval) : 372 | (forall k, In k rem -> ~In_dom k al) -> 373 | remove_assets al rem = al. 374 | induction al as [|[h [obl u]] ar IH]. 375 | - intros _. reflexivity. 376 | - intros H1. simpl. destruct (in_dec hashval_eq_dec h rem) as [D1|D1]. 377 | + exfalso. apply (H1 h D1). now left. 378 | + f_equal. apply IH. intros k H2 H3. apply (H1 k H2). now right. 379 | Qed. 380 | 381 | Lemma fnl_remove_assets (al:list asset) (rem:list hashval) : 382 | fnl al -> fnl (remove_assets al rem). 383 | intros H. induction H as [|h u al H1 H2 IH]. 384 | - apply fnl_N. 385 | - simpl. destruct (in_dec hashval_eq_dec h rem) as [H3|H3]. 386 | + exact IH. 387 | + apply fnl_C. 388 | * intros H4. apply H1. apply In_In_dom_lem in H4. 389 | destruct H4 as [v H4]. 390 | apply remove_assets_iff in H4. destruct H4 as [H5 H6]. 391 | apply In_In_dom_lem. exists v. assumption. 392 | * exact IH. 393 | Qed. 394 | 395 | Definition hashassetlist (al:list asset) : option hashval := ohashlist (map hashasset al). 396 | 397 | Lemma hashassetlistinj al bl : hashassetlist al = hashassetlist bl -> al = bl. 398 | intros H1. 399 | change (ohashlist (map hashasset al) = ohashlist (map hashasset bl)) in H1. 400 | apply ohashlistinj in H1. 401 | revert bl H1. induction al as [|a al IH]; intros [|b bl] H1. 402 | - reflexivity. 403 | - discriminate H1. 404 | - discriminate H1. 405 | - simpl in H1. inversion H1. apply hashassetinj in H0. subst b. f_equal. 406 | now apply IH. 407 | Qed. 408 | 409 | Definition asset_value_sum (al:list asset) : nat := 410 | fold_right plus 0 (map (fun a => match asset_value a with Some u => u | None => 0 end) al). 411 | 412 | Lemma asset_value_sum_value_nil : 413 | asset_value_sum nil = 0. 414 | reflexivity. 415 | Qed. 416 | 417 | Lemma asset_value_sum_value_cons a al u : 418 | asset_value a = Some u -> 419 | asset_value_sum (a::al) = u + asset_value_sum al. 420 | intros H1. unfold asset_value_sum. simpl. rewrite H1. 421 | reflexivity. 422 | Qed. 423 | 424 | Lemma asset_value_sum_novalue_cons a al : 425 | asset_value a = None -> 426 | asset_value_sum (a::al) = asset_value_sum al. 427 | intros H1. unfold asset_value_sum. simpl. rewrite H1. 428 | reflexivity. 429 | Qed. 430 | 431 | Lemma asset_value_sum_app (l r:list asset) : 432 | asset_value_sum l 433 | + 434 | asset_value_sum r 435 | = 436 | asset_value_sum (l ++ r). 437 | unfold asset_value_sum. induction l as [|[h u] l IH]. 438 | - simpl. reflexivity. 439 | - simpl. rewrite <- plus_assoc. rewrite IH. reflexivity. 440 | Qed. 441 | 442 | Lemma app_perm_asset_value_sum (l r:list asset) : 443 | app_perm l r -> 444 | asset_value_sum l = asset_value_sum r. 445 | intros H. induction H as [l r|l1 r1 l2 r2 H1 IH1 H2 IH2|l|l r H1 IH1|l r s H1 IH1 H2 IH2]. 446 | - rewrite <- asset_value_sum_app. rewrite <- asset_value_sum_app. omega. 447 | - rewrite <- asset_value_sum_app. rewrite <- asset_value_sum_app. omega. 448 | - omega. 449 | - omega. 450 | - omega. 451 | Qed. 452 | 453 | Lemma remove_assets_app (al bl:list asset) (rl:list hashval) : 454 | remove_assets (al ++ bl) rl = (remove_assets al rl) ++ (remove_assets bl) rl. 455 | induction al as [|[h u] al IH]. 456 | - reflexivity. 457 | - simpl. destruct (in_dec hashval_eq_dec h rl) as [H1|H1]. 458 | + exact IH. 459 | + rewrite IH. reflexivity. 460 | Qed. 461 | 462 | Lemma remove_assets_nil (l:list asset) : 463 | remove_assets l nil = l. 464 | induction l as [|[h u] l IH]. 465 | - reflexivity. 466 | - simpl. rewrite IH. reflexivity. 467 | Qed. 468 | 469 | Lemma remove_assets_nIn_cons (l:list asset) (h:hashval) (rem:list hashval) : 470 | (~exists u, In (h,u) l) -> 471 | remove_assets l (h::rem) = remove_assets l rem. 472 | induction l as [|[k v] l IH]; intros H1. 473 | - reflexivity. 474 | - simpl. destruct (hashval_eq_dec h k) as [D1|D1]. 475 | + exfalso. apply H1. exists v. left. congruence. 476 | + destruct (in_dec hashval_eq_dec k rem) as [D2|D2]. 477 | * apply IH. intros [u H2]. apply H1. exists u. now right. 478 | * f_equal. apply IH. intros [u H2]. apply H1. exists u. now right. 479 | Qed. 480 | 481 | Lemma remove_assets_app2 (l:list asset) (rem1 rem2:list hashval) : 482 | remove_assets l (rem1 ++ rem2) = remove_assets (remove_assets l rem1) rem2. 483 | induction l as [ |[k v] l IH]. 484 | - reflexivity. 485 | - simpl. destruct (in_dec hashval_eq_dec k (rem1 ++ rem2)) as [D1|D1]. 486 | + destruct (in_dec hashval_eq_dec k rem1) as [D2|D2]. 487 | * exact IH. 488 | * { simpl. destruct (in_dec hashval_eq_dec k rem2) as [D3|D3]. 489 | - exact IH. 490 | - exfalso. apply in_app_or in D1. tauto. 491 | } 492 | + destruct (in_dec hashval_eq_dec k rem1) as [D2|D2]. 493 | * exfalso. apply D1. apply in_or_app. tauto. 494 | * { simpl. destruct (in_dec hashval_eq_dec k rem2) as [D3|D3]. 495 | - exfalso. apply D1. apply in_or_app. tauto. 496 | - f_equal. exact IH. 497 | } 498 | Qed. 499 | 500 | Lemma remove_assets_eq (l:list asset) (rl1 rl2:list hashval) : 501 | (forall h, In_dom h l -> (In h rl1 <-> In h rl2)) -> 502 | remove_assets l rl1 = remove_assets l rl2. 503 | induction l as [|[h u] l IH]. 504 | - intros _. reflexivity. 505 | - intros H1. simpl. 506 | destruct (in_dec hashval_eq_dec h rl1) as [H2|H2]; 507 | destruct (in_dec hashval_eq_dec h rl2) as [H3|H3]. 508 | + apply IH. intros k H4. apply H1. now right. 509 | + exfalso. apply H3. apply H1. 510 | * now left. 511 | * assumption. 512 | + exfalso. apply H2. apply H1. 513 | * now left. 514 | * assumption. 515 | + f_equal. apply IH. intros k H4. apply H1. now right. 516 | Qed. 517 | 518 | Lemma asset_value_sum_asset_shift (al:list asset) h a u rl : 519 | fnl al -> 520 | no_dups (h::rl) -> 521 | In (h,a) al -> 522 | asset_value (h,a) = Some u -> 523 | u + asset_value_sum (remove_assets al (h::rl)) = asset_value_sum (remove_assets al rl). 524 | intros H H0. induction H as [|k v al H1 H2 IH]. 525 | - simpl. tauto. 526 | - simpl. intros H3 E0. destruct (hashval_eq_dec h k) as [H4|H4]. 527 | + destruct (in_dec hashval_eq_dec k rl) as [H5|H5]. 528 | * exfalso. inversion H0. subst k. contradiction. 529 | * { destruct H3 as [H3|H3]. 530 | - inversion H3. 531 | rewrite (asset_value_sum_value_cons _ _ _ E0). 532 | f_equal. 533 | subst k. 534 | apply app_perm_asset_value_sum. 535 | rewrite (remove_assets_eq al (h::rl) rl). 536 | + apply app_perm_ref. 537 | + intros k H8. split. 538 | * { intros [H9|H9]. 539 | - exfalso. apply H1. subst k. assumption. 540 | - assumption. 541 | } 542 | * intros H9. now right. 543 | - exfalso. apply H1. apply In_In_dom_lem. exists a. subst k. assumption. 544 | } 545 | + destruct (in_dec hashval_eq_dec k rl) as [H5|H5]. 546 | * { destruct H3 as [H3|H3]. 547 | - inversion H3. exfalso. congruence. 548 | - now apply IH. 549 | } 550 | * { destruct H3 as [H3|H3]. 551 | - inversion H3. exfalso. congruence. 552 | - destruct (asset_value (k,v)) as [w|] eqn: E1. 553 | + rewrite (asset_value_sum_value_cons _ _ _ E1). 554 | rewrite (asset_value_sum_value_cons _ _ _ E1). 555 | apply IH in H3. 556 | * omega. 557 | * exact E0. 558 | + rewrite (asset_value_sum_novalue_cons _ _ E1). 559 | rewrite (asset_value_sum_novalue_cons _ _ E1). 560 | exact (IH H3 E0). 561 | } 562 | Qed. 563 | 564 | Lemma asset_value_sum_asset_shift_non (al:list asset) h a rl : 565 | fnl al -> 566 | no_dups (h::rl) -> 567 | In (h,a) al -> 568 | asset_value (h,a) = None -> 569 | asset_value_sum (remove_assets al (h::rl)) = asset_value_sum (remove_assets al rl). 570 | intros H H0. induction H as [|k v al H1 H2 IH]. 571 | - simpl. tauto. 572 | - simpl. intros H3 H3'. destruct (hashval_eq_dec h k) as [H4|H4]. 573 | + destruct (in_dec hashval_eq_dec k rl) as [H5|H5]. 574 | * exfalso. inversion H0. subst k. contradiction. 575 | * { destruct H3 as [H3|H3]. 576 | - inversion H3. 577 | rewrite (asset_value_sum_novalue_cons _ _ H3'). 578 | change (asset_value_sum (remove_assets al (h :: rl)) = asset_value_sum (remove_assets al rl)). 579 | subst k. 580 | apply app_perm_asset_value_sum. 581 | rewrite (remove_assets_eq al (h::rl) rl). 582 | + apply app_perm_ref. 583 | + intros k H8. split. 584 | * { intros [H9|H9]. 585 | - exfalso. apply H1. subst k. assumption. 586 | - assumption. 587 | } 588 | * intros H9. now right. 589 | - exfalso. apply H1. apply In_In_dom_lem. exists a. subst k. assumption. 590 | } 591 | + destruct (in_dec hashval_eq_dec k rl) as [H5|H5]. 592 | * { destruct H3 as [H3|H3]. 593 | - inversion H3. exfalso. congruence. 594 | - now apply IH. 595 | } 596 | * { destruct H3 as [H3|H3]. 597 | - inversion H3. exfalso. congruence. 598 | - destruct (asset_value (k,v)) as [w|] eqn: E1. 599 | + rewrite (asset_value_sum_value_cons _ _ _ E1). 600 | rewrite (asset_value_sum_value_cons _ _ _ E1). 601 | apply IH in H3. 602 | * omega. 603 | * exact H3'. 604 | + rewrite (asset_value_sum_novalue_cons _ _ E1). 605 | rewrite (asset_value_sum_novalue_cons _ _ E1). 606 | exact (IH H3 H3'). 607 | } 608 | Qed. 609 | 610 | Lemma add_vout_lem (txh:hashval) (outpl:list addr_preasset) (i:nat) alpha h u : 611 | In (alpha,(h,u)) (add_vout txh outpl i) 612 | <-> 613 | exists j, nth_error outpl j = value (alpha,u) 614 | /\ h = hashpair txh (hashnat (i + j)). 615 | revert i. induction outpl as [|[k [obl v]] outpr IH]; intros i. 616 | - simpl. split. 617 | + tauto. 618 | + intros [[|j] [H1 _]]; discriminate H1. 619 | - simpl. split. 620 | + intros [H1|H1]. 621 | * { inversion H1. exists 0. split. 622 | - reflexivity. 623 | - f_equal. f_equal. omega. 624 | } 625 | * { apply IH in H1. destruct H1 as [j [H2 H3]]. exists (S j). split. 626 | - exact H2. 627 | - rewrite H3. f_equal. f_equal. omega. 628 | } 629 | + intros [[|j] [H1 H2]]. 630 | * left. simpl in H1. inversion H1. rewrite H2. 631 | f_equal. f_equal. f_equal. f_equal. omega. 632 | * { right. apply IH. exists j. split. 633 | - exact H1. 634 | - rewrite H2. f_equal. f_equal. omega. 635 | } 636 | Qed. 637 | -------------------------------------------------------------------------------- /Blocks.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** Blocks: Blocks are given by headers and deltas. 6 | We define types of block chains and block header chains. 7 | We leave the reward function, how 'hits' are checked and how targets are calculated as parameters. 8 | The block contains a ctree corresponding to sufficient information about the previous ledger 9 | to process the block (ctree_of_Block) obtained by grafting the information in the block delta 10 | to the information in the block header. All the transactions in the block (including the 11 | coinstake in the header) can be combined to give one large transaction (tx_of_Block). 12 | If the block is valid (valid_Block), then the combined transaction will be valid and supported. 13 | If we start from a valid genesis ledger and a block chain is built from valid blocks (valid_BlockChain), 14 | then we are guaranteed that the ending ledger is also valid (lastledgerroot_valid). 15 | Furthermore, we know precisely how many currency units are in the ledger as of 16 | a certain block height (lastledgerroot_sum_correct). 17 | **) 18 | 19 | Require Export CTreeGrafting. 20 | 21 | (*** Idea: keep up with how fast the past N blocks and use this to determine the current target. I'll use 6 blocks just as a concrete option here. ***) 22 | Definition targetinfo : Type := prod nat (list nat). 23 | Definition nexttargetinfo (ti : targetinfo) (tm : nat) : targetinfo := 24 | match ti with 25 | | (prevtm,nil) => (tm,(tm - prevtm)::nil) 26 | | (prevtm,c1::nil) => (tm,c1::(tm - prevtm)::nil) 27 | | (prevtm,c2::c1::nil) => (tm,c2::c1::(tm - prevtm)::nil) 28 | | (prevtm,c3::c2::c1::nil) => (tm,c3::c2::c1::(tm - prevtm)::nil) 29 | | (prevtm,c4::c3::c2::c1::nil) => (tm,c4::c3::c2::c1::(tm - prevtm)::nil) 30 | | (prevtm,c5::c4::c3::c2::c1::nil) => (tm,c5::c4::c3::c2::c1::(tm - prevtm)::nil) 31 | | (prevtm,c6::c5::c4::c3::c2::c1::_) => (tm,c5::c4::c3::c2::c1::(tm - prevtm)::nil) 32 | end. 33 | 34 | Definition hashtargetinfo (ti:targetinfo) : hashval := 35 | let (n,ml) := ti in 36 | hashpair (hashnat n) (hashlist (map hashnat ml)). 37 | 38 | Definition targetinfo_timestamp (ti:targetinfo) : nat := fst ti. 39 | 40 | Record BlockHeader : Type := 41 | mkBlockHeader 42 | { 43 | prevblockhash : option hashval; 44 | newledgerroot : hashval; 45 | targetinfohash : hashval; 46 | timestamp : nat; 47 | stake : nat; 48 | stakeaddr : addr; 49 | stakeassetid : hashval; 50 | blocksignat : signat; 51 | prevledger : ctree 160 52 | }. 53 | 54 | Record BlockDelta : Type := 55 | mkBlockDelta 56 | { 57 | totalfees : nat; 58 | stakeoutput : list addr_preasset; 59 | prevledgergraft : cgraft; 60 | blockdelta_stxl : list sTx 61 | }. 62 | 63 | Definition Block : Type := 64 | prod BlockHeader BlockDelta. 65 | 66 | Definition coinstake (b:Block) : Tx := 67 | let (bh,bd) := b in 68 | ((stakeaddr bh,stakeassetid bh)::nil,stakeoutput bd). 69 | 70 | Definition hitfun : Type := option hashval -> nat -> nat -> nat -> addr -> Prop. 71 | 72 | Definition targetfun : Type := targetinfo -> nat. 73 | 74 | Definition hash_BlockHeader (bh:BlockHeader) : hashval := 75 | hashopair2 (prevblockhash bh) (hashlist (newledgerroot bh::targetinfohash bh::hashnat (timestamp bh)::hashnat (stake bh)::hashaddr (stakeaddr bh)::(stakeassetid bh)::nil)). 76 | 77 | (*** A currency asset can be used to stake as long as it will not mature in the next 1000 blocks. (1000 is arbitrary, of course.) ***) 78 | Definition not_close_to_mature : nat := 1000. 79 | 80 | (*** The output from staking cannot be spent until 1000 blocks later. Again, 1000 is arbitrary here. ***) 81 | Definition maturation_post_staking : nat := 1000. 82 | 83 | Definition valid_BlockHeader (blockheight:nat) (check_hit : hitfun) (targetf : targetfun) (plr:hashval) (ti:targetinfo) (bh:BlockHeader) : Prop := 84 | (timestamp bh > targetinfo_timestamp ti /\ hashtargetinfo ti = targetinfohash bh) 85 | /\ 86 | check_signat (hash_BlockHeader bh) (blocksignat bh) (stakeaddr bh) 87 | /\ 88 | check_hit (prevblockhash bh) (timestamp bh) (targetf ti) (stake bh) (stakeaddr bh) 89 | /\ 90 | ctree_hashroot (prevledger bh) = plr 91 | /\ 92 | (exists al m n, 93 | ctree_supports_asset (stakeassetid bh,((al,(m,n)),currency (stake bh))) (prevledger bh) (stakeaddr bh) 94 | /\ 95 | n > blockheight + not_close_to_mature) 96 | . 97 | 98 | Fixpoint sTxs_allinputs (stxl : list sTx) : list addr_assetid := 99 | match stxl with 100 | | ((inpl,_),s)::stxr => inpl ++ sTxs_allinputs stxr 101 | | nil => nil 102 | end. 103 | 104 | Fixpoint sTxs_alloutputs (stxl : list sTx) : list addr_preasset := 105 | match stxl with 106 | | ((_,outpl),s)::stxr => outpl ++ sTxs_alloutputs stxr 107 | | nil => nil 108 | end. 109 | 110 | Lemma sTxs_allinputs_iff stxl alpha h : 111 | In (alpha,h) (sTxs_allinputs stxl) <-> exists tx, In_dom tx stxl /\ In (alpha,h) (tx_inputs tx). 112 | induction stxl as [|[[inpl outpl] s] stxr IH]; split. 113 | - simpl. tauto. 114 | - intros [tx [H1 _]]. contradiction H1. 115 | - simpl. intros H1. apply in_app_or in H1. destruct H1 as [H1|H1]. 116 | + exists (inpl,outpl). simpl. tauto. 117 | + apply IH in H1. destruct H1 as [tx H2]. exists tx. tauto. 118 | - intros [tx [[H1|H1] H2]]. 119 | + simpl. apply in_or_app. left. rewrite H1 in H2. exact H2. 120 | + simpl. apply in_or_app. right. apply IH. exists tx. tauto. 121 | Qed. 122 | 123 | Lemma sTxs_alloutputs_iff stxl alpha u : 124 | In (alpha,u) (sTxs_alloutputs stxl) <-> exists tx, In_dom tx stxl /\ In (alpha,u) (tx_outputs tx). 125 | induction stxl as [|[[inpl outpl] s] stxr IH]; split. 126 | - simpl. tauto. 127 | - intros [tx [H1 _]]. contradiction H1. 128 | - simpl. intros H1. apply in_app_or in H1. destruct H1 as [H1|H1]. 129 | + exists (inpl,outpl). simpl. tauto. 130 | + apply IH in H1. destruct H1 as [tx H2]. exists tx. tauto. 131 | - intros [tx [[H1|H1] H2]]. 132 | + simpl. apply in_or_app. left. rewrite H1 in H2. exact H2. 133 | + simpl. apply in_or_app. right. apply IH. exists tx. tauto. 134 | Qed. 135 | 136 | Definition ctree_of_Header (b:Block) : ctree 160 := 137 | let (bh,bd) := b in 138 | prevledger bh. 139 | 140 | Definition ctree_of_Block (b:Block) : ctree 160 := 141 | let (bh,bd) := b in 142 | ctree_cgraft (prevledgergraft bd) (prevledger bh). 143 | 144 | Definition timestamp_of_Block (b:Block) : nat := 145 | let (bh,bd) := b in 146 | timestamp bh. 147 | 148 | (*** One Tx combining all the Txs in the block, including the coinstake. ***) 149 | Definition tx_of_Block (b:Block) : Tx := 150 | let (bh,bd) := b in 151 | ((stakeaddr bh,stakeassetid bh)::sTxs_allinputs (blockdelta_stxl bd),stakeoutput bd ++ sTxs_alloutputs (blockdelta_stxl bd)). 152 | 153 | Definition valid_Block (blockheight:nat) (rew:nat) (check_hit : hitfun) (targetf : targetfun) plr ti (b:Block) : Prop := 154 | let (bh,bd) := b in 155 | (*** The header is valid. ***) 156 | valid_BlockHeader blockheight check_hit targetf plr ti bh 157 | /\ 158 | (*** The asset staked must be one of the outputs. This is especially important if it's been loaned for staking. 159 | The other stake outputs must explicitly say that they cannot be spent until at least blockheight + maturation_post_staking. ***) 160 | (exists obl:obligation, exists i:nat, 161 | ctree_supports_asset (stakeassetid bh,(obl,currency (stake bh))) (prevledger bh) (stakeaddr bh) 162 | /\ 163 | nth_error (stakeoutput bd) i = value (stakeaddr bh,(obl,currency (stake bh))) 164 | /\ 165 | (forall alpha al m n u j, j <> i /\ nth_error (stakeoutput bd) j = value (alpha,((al,(m,n)),u)) -> n > blockheight + maturation_post_staking)) 166 | /\ 167 | (*** The addresses of the coinstake output are supported by the grafted ctree. ***) 168 | ((forall alpha ou, In (alpha,ou) (stakeoutput bd) -> ctree_supports_addr (ctree_of_Block (bh,bd)) alpha) 169 | /\ 170 | (*** and the total output values match the declared stake and reward and fees ***) 171 | (asset_value_out (stakeoutput bd) = stake bh + rew + totalfees bd)) 172 | /\ 173 | (*** There are no duplicate transactions. ***) 174 | no_dups (blockdelta_stxl bd) 175 | /\ 176 | (*** Each transaction in the delta has supported elaborated assets and has valid signatures. ***) 177 | (forall tx sl, In (tx,sl) (blockdelta_stxl bd) -> 178 | exists einpl:list asset, 179 | tx_signatures_valid blockheight einpl (tx,sl) 180 | /\ 181 | (forall beta h, In (beta,h) (tx_inputs tx) -> exists a, In a einpl /\ assetid a = h /\ ctree_supports_asset a (ctree_of_Block (bh,bd)) beta)) 182 | /\ 183 | (*** Each transaction in the delta is valid. ***) 184 | (forall stx, In stx (blockdelta_stxl bd) -> tx_valid (fst stx)) 185 | /\ 186 | (*** No transaction has the stake asset as an input. ***) 187 | (forall tx, In_dom tx (blockdelta_stxl bd) -> ~ In (stakeaddr bh,stakeassetid bh) (tx_inputs tx)) 188 | /\ 189 | (*** No distinct transactions try to spend the same asset. ***) 190 | (forall tx1 sl1 tx2 sl2 alpha h, 191 | In (tx1,sl1) (blockdelta_stxl bd) -> 192 | In (tx2,sl2) (blockdelta_stxl bd) -> 193 | In (alpha,h) (tx_inputs tx1) -> 194 | In (alpha,h) (tx_inputs tx2) -> 195 | tx1 = tx2 /\ sl1 = sl2) 196 | /\ 197 | (*** The cgraft is valid. ***) 198 | cgraft_valid (prevledgergraft bd) 199 | /\ 200 | (*** Every transaction is supported by the grafted ctree with 0 reward. ***) 201 | (forall tx, In_dom tx (blockdelta_stxl bd) 202 | -> exists fee, ctree_supports_tx tx (ctree_of_Block (bh,bd)) fee 0) 203 | (*** The total inputs and outputs match up with the declared fee. ***) 204 | /\ 205 | (exists utot:nat, 206 | ctree_asset_value_in (ctree_of_Block (bh,bd)) (sTxs_allinputs (blockdelta_stxl bd)) utot 207 | /\ 208 | asset_value_out (sTxs_alloutputs (blockdelta_stxl bd)) + (totalfees bd) = utot) 209 | /\ 210 | exists T:ctree 160, 211 | ctree_hashroot T = newledgerroot bh 212 | /\ tx_octree_trans (tx_of_Block (bh,bd)) (Some (ctree_of_Block (bh,bd))) = Some T. 213 | 214 | Inductive BlockChain : nat -> Type := 215 | | BlockChainGen : Block -> BlockChain 0 216 | | BlockChainSucc n : BlockChain n -> Block -> BlockChain (S n). 217 | 218 | Inductive BlockHeaderChain : nat -> Type := 219 | | BlockHeaderChainGen : BlockHeader -> BlockHeaderChain 0 220 | | BlockHeaderChainSucc n : BlockHeaderChain n -> BlockHeader -> BlockHeaderChain (S n). 221 | 222 | Fixpoint BlockChain_headers {n} (bc : BlockChain n) : BlockHeaderChain n := 223 | match bc with 224 | | BlockChainGen (bh,_) => BlockHeaderChainGen bh 225 | | BlockChainSucc n bc (bh,_) => BlockHeaderChainSucc n (BlockChain_headers bc) bh 226 | end. 227 | 228 | Definition lastledgerroot_of_BlockChain {n} (bc:BlockChain n) : hashval := 229 | match bc with 230 | | BlockChainGen (bh,_) => newledgerroot bh 231 | | BlockChainSucc n bc (bh,_) => newledgerroot bh 232 | end. 233 | 234 | Definition lastledgerroot_of_BlockHeaderChain {n} (bc:BlockHeaderChain n) : hashval := 235 | match bc with 236 | | BlockHeaderChainGen bh => newledgerroot bh 237 | | BlockHeaderChainSucc n bc bh => newledgerroot bh 238 | end. 239 | 240 | Fixpoint valid_BlockChain (rewfn:nat -> nat) (check_hit:hitfun) (targetf:targetfun) genlr genti ti {n} (bc : BlockChain n) : Prop := 241 | match bc with 242 | | BlockChainGen b => valid_Block 0 (rewfn 0) check_hit targetf genlr genti b /\ ti = nexttargetinfo genti (timestamp_of_Block b) 243 | | BlockChainSucc n bc b => exists ti', valid_BlockChain rewfn check_hit targetf genlr genti ti' bc /\ valid_Block (S n) (rewfn (S n)) check_hit targetf (lastledgerroot_of_BlockChain bc) ti' b /\ ti = nexttargetinfo ti' (timestamp_of_Block b) 244 | end. 245 | 246 | Fixpoint valid_BlockHeaderChain (check_hit:hitfun) (targetf:targetfun) genlr genti ti {n} (bc : BlockHeaderChain n) : Prop := 247 | match bc with 248 | | BlockHeaderChainGen bh => valid_BlockHeader 0 check_hit targetf genlr genti bh /\ ti = nexttargetinfo genti (timestamp bh) 249 | | BlockHeaderChainSucc n bc bh => exists ti', valid_BlockHeaderChain check_hit targetf genlr genti ti' bc /\ valid_BlockHeader (S n) check_hit targetf (lastledgerroot_of_BlockHeaderChain bc) ti' bh /\ ti = nexttargetinfo ti' (timestamp bh) 250 | end. 251 | 252 | Lemma tx_of_Block_valid blockheight rew check_hit targetf plr ti (b:Block) : 253 | valid_Block blockheight rew check_hit targetf plr ti b -> 254 | tx_valid (tx_of_Block b). 255 | destruct b as [bh bd]. 256 | intros HvB. generalize HvB. 257 | intros [[HvBaa [HvBab [HvBac [HvBad [HvBae HvBaf]]]]] [HvBag [HvBah [HvBb [HvBc [HvBd [HvBe [HvBf [HvBg [HvBh [HvBi HvBj]]]]]]]]]]]. 258 | split. 259 | - simpl. discriminate. 260 | - simpl. apply no_dups_cons. 261 | + intros H1. apply sTxs_allinputs_iff in H1. 262 | destruct H1 as [[inpl outpl] [H2 H3]]. 263 | revert H2 H3. apply HvBe. 264 | + revert HvBb HvBd HvBf. unfold sTxs_allinputs. generalize (blockdelta_stxl bd). 265 | clear. 266 | (*** This could be made into a reasonable lemma instead of giving a double induction in a subproof. ***) 267 | induction l as [|[[inpl outpl] sl] txl IH]. 268 | * simpl. intros H0 H1 H2. apply no_dups_nil. 269 | * { simpl. intros H0 H1 H2. apply no_dups_app. 270 | - assert (L1: (inpl, outpl, sl) = (inpl, outpl, sl) \/ In (inpl, outpl, sl) txl) by now left. 271 | destruct (H1 ((inpl,outpl),sl) L1) as [_ H1b]. 272 | exact H1b. 273 | - apply IH. 274 | + inversion H0. assumption. 275 | + intros stx H3. apply H1. now right. 276 | + intros tx1 sl1 tx2 sl2 alpha h H3 H4. apply H2; now right. 277 | - revert H0 H2. clear. 278 | induction txl as [|[[inpl2 outpl2] sl2] txr IH]; intros H0 H2 [alpha h] H3. 279 | + simpl. tauto. 280 | + intros H4. apply in_app_or in H4. 281 | destruct H4 as [H4|H4]. 282 | * inversion H0. apply H5. left. 283 | assert (L1: (inpl,outpl) = (inpl2,outpl2) /\ sl = sl2). 284 | { apply (H2 (inpl,outpl) sl (inpl2,outpl2) sl2 alpha h). 285 | - now left. 286 | - right. now left. 287 | - exact H3. 288 | - exact H4. 289 | } 290 | destruct L1 as [L1a L1b]. 291 | rewrite L1a. rewrite L1b. 292 | reflexivity. 293 | * { revert H4. apply IH. 294 | - inversion H0. apply no_dups_cons. 295 | + intros H6. apply H4. now right. 296 | + inversion H5. assumption. 297 | - intros tx3 sl3 tx4 sl4 beta k H6 H7. apply H2. 298 | + simpl. tauto. 299 | + simpl. tauto. 300 | - exact H3. 301 | } 302 | } 303 | Qed. 304 | 305 | Lemma tx_of_Block_input_iff alpha h (b : Block) : 306 | In (alpha,h) (tx_inputs (tx_of_Block b)) 307 | <-> 308 | ((alpha = stakeaddr (fst b) /\ h = stakeassetid (fst b)) 309 | \/ 310 | exists tx sl, In (tx,sl) (blockdelta_stxl (snd b)) /\ In (alpha,h) (tx_inputs tx)). 311 | destruct b as [bh bd]. simpl. split; intros [H1|H1]. 312 | - left. inversion H1. tauto. 313 | - right. apply sTxs_allinputs_iff in H1. 314 | destruct H1 as [tx [H1 H2]]. apply In_In_dom_lem in H1. 315 | destruct H1 as [sl H1]. exists tx. exists sl. tauto. 316 | - left. destruct H1 as [H2 H3]. congruence. 317 | - right. apply sTxs_allinputs_iff. destruct H1 as [tx [sl [H1 H2]]]. 318 | exists tx. split. 319 | + revert H1. apply In_In_dom_lem_2. 320 | + exact H2. 321 | Qed. 322 | 323 | Lemma tx_of_Block_output_iff alpha obju (b : Block) : 324 | In (alpha, obju) (tx_outputs (tx_of_Block b)) 325 | <-> 326 | (In (alpha, obju) (stakeoutput (snd b)) 327 | \/ 328 | exists tx sl, In (tx,sl) (blockdelta_stxl (snd b)) /\ In (alpha,obju) (tx_outputs tx)). 329 | destruct b as [bh bd]. simpl. split; intros H1. 330 | - apply in_app_or in H1. destruct H1 as [H1|H1]. 331 | + now left. 332 | + right. apply sTxs_alloutputs_iff in H1. 333 | destruct H1 as [tx [H1 H2]]. apply In_In_dom_lem in H1. 334 | destruct H1 as [sl H1]. exists tx. exists sl. tauto. 335 | - apply in_or_app. destruct H1 as [H1|H1]. 336 | + now left. 337 | + right. apply sTxs_alloutputs_iff. destruct H1 as [tx [sl [H1 H2]]]. 338 | exists tx. split. 339 | * revert H1. apply In_In_dom_lem_2. 340 | * exact H2. 341 | Qed. 342 | 343 | Opaque ctree_supports_addr. 344 | 345 | (*** This lemma was pulled out from the main proof to help Coq process the main proof. ***) 346 | Lemma tx_of_Block_supported_lem_1 blockheight rew (bh:BlockHeader) (bd:BlockDelta) utot : 347 | ctree_valid (ctree_of_Block (bh, bd)) -> 348 | subqc (prevledger bh) (ctree_of_Block (bh, bd)) -> 349 | (exists (al : list addr) (m n : nat), 350 | ctree_supports_asset 351 | (stakeassetid bh, (al, (m, n), currency (stake bh))) 352 | (prevledger bh) (stakeaddr bh) /\ 353 | n > blockheight + not_close_to_mature) -> 354 | asset_value_out (stakeoutput bd) = stake bh + rew + totalfees bd -> 355 | asset_value_out (sTxs_alloutputs (blockdelta_stxl bd)) + totalfees bd = utot -> 356 | asset_value_out (tx_outputs (tx_of_Block (bh, bd))) + 0 = 357 | stake bh + utot + rew. 358 | intros HT Lsubqc HvBae H5 H3. 359 | simpl. rewrite asset_value_out_app. 360 | omega. 361 | Qed. 362 | 363 | Theorem tx_of_Block_supported blockheight rew check_hit targetf plr ti (b:Block) : 364 | ctree_valid (ctree_of_Block b) -> 365 | valid_Block blockheight rew check_hit targetf plr ti b -> 366 | ctree_supports_tx (tx_of_Block b) (ctree_of_Block b) 0 rew. 367 | destruct b as [bh bd]. intros HT HvB. generalize HvB. 368 | intros [[HvBaa [HvBab [HvBac [HvBad HvBae]]]] [HvBaf [HvBag [HvBb [HvBc [HvBd [HvBe [HvBf [HvBg [HvBh [HvBi HvBj]]]]]]]]]]]. 369 | generalize HvBag. intros [HvBag1 HvBag2]. 370 | assert (Lsubqc: subqc (prevledger bh) (ctree_of_Block (bh,bd))). 371 | { unfold ctree_of_Block. apply ctree_cgraft_subqc. exact HvBg. } 372 | split. 373 | - intros alpha obju H1. 374 | apply tx_of_Block_output_iff in H1. 375 | destruct H1 as [H1|[tx' [sl' [H2 H3]]]]. 376 | + exact (HvBag1 alpha obju H1). 377 | + apply In_In_dom_lem_2 in H2. 378 | destruct (HvBh tx' H2) as [fee [H2a _]]. 379 | revert H3. apply H2a. 380 | - destruct HvBi as [utot [H2 H3]]. 381 | exists (stake bh + utot). 382 | split. 383 | + change (ctree_asset_value_in (ctree_of_Block (bh, bd)) 384 | ((stakeaddr bh, stakeassetid bh) 385 | :: sTxs_allinputs (blockdelta_stxl bd)) 386 | (stake bh + utot)). 387 | destruct HvBae as [al [m [n [H4 H5]]]]. 388 | apply ctree_asset_value_in_cons with (a := (stakeassetid bh,((al,(m,n)),currency (stake bh)))). 389 | * exact H2. 390 | * revert H4. apply subqc_supports_asset. exact Lsubqc. 391 | * reflexivity. 392 | * reflexivity. 393 | + revert HT Lsubqc HvBae HvBag2 H3. 394 | apply tx_of_Block_supported_lem_1. 395 | Qed. 396 | 397 | Definition ledgerroot_valid (h:hashval) : Prop := 398 | ctree_valid (ctreeH 160 h). 399 | 400 | Opaque tx_octree_trans. 401 | 402 | Lemma ctree_hashroot_ctree_H {n} (h:hashval) : 403 | ctree_hashroot (ctreeH n h) = h. 404 | destruct n; simpl; reflexivity. 405 | Qed. 406 | 407 | Opaque mtree_approx_fun_p. 408 | Opaque ctree_mtree. 409 | Opaque tx_of_Block. 410 | Opaque mtree_hashroot. 411 | Opaque ctree_hashroot. 412 | Opaque ctree_cgraft. 413 | Opaque cgraft_assoc. 414 | Opaque subqm. 415 | 416 | Lemma endledgerroot_plr_valid_lem1 lroot (bh:BlockHeader) bd (f:statefun) : 417 | mtree_approx_fun_p (ctree_mtree (ctreeH 160 lroot)) f -> 418 | ctree_hashroot (prevledger bh) = lroot -> 419 | cgraft_valid (prevledgergraft bd) -> 420 | mtree_approx_fun_p (ctree_mtree (ctree_of_Block (bh, bd))) f. 421 | intros H1 H2 H3. 422 | apply (mtree_hashroot_mtree_approx_fun_p (ctree_mtree (ctreeH 160 lroot))). 423 | - change (mtree_hashroot (ctree_mtree (ctreeH 160 lroot)) = 424 | mtree_hashroot (ctree_mtree (ctree_of_Block (bh, bd)))). 425 | rewrite mtree_hashroot_ctree_hashroot. 426 | rewrite mtree_hashroot_ctree_hashroot. 427 | rewrite ctree_hashroot_ctree_H. 428 | unfold ctree_of_Block. 429 | assert (L1: subqc (prevledger bh) (ctree_cgraft (prevledgergraft bd) (prevledger bh))). 430 | { exact (ctree_cgraft_subqc (prevledgergraft bd) (prevledger bh) H3). } 431 | unfold subqc in L1. 432 | apply subqm_hashroot_eq in L1. 433 | rewrite mtree_hashroot_ctree_hashroot in L1. 434 | rewrite mtree_hashroot_ctree_hashroot in L1. 435 | rewrite H2 in L1. 436 | exact L1. 437 | - exact H1. 438 | Qed. 439 | 440 | Theorem lastledgerroot_valid rewfn check_hit targetf genlr genti ti {n} (bc : BlockChain n) : 441 | ledgerroot_valid genlr -> 442 | valid_BlockChain rewfn check_hit targetf genlr genti ti bc -> 443 | ledgerroot_valid (lastledgerroot_of_BlockChain bc). 444 | intros H1. revert ti. induction bc as [[bh bd]|n bc IH [bh bd]]; intros ti. 445 | - intros H2. generalize H2. intros [[[H2aa [H2ab [H2ac [H2ad H2ae]]]] [H2af [H2ag [H2b [H2c [H2d [H2e [H2f [H2g [H2h [H2i H2j]]]]]]]]]]] Hti]. 446 | assert (LT: ctree_valid (ctree_of_Block (bh,bd))). 447 | { unfold ctree_of_Block. apply ctree_valid_cgraft_valid. 448 | - revert H1. unfold ledgerroot_valid. unfold ctree_valid. 449 | apply mtree_hashroot_eq_valid. 450 | rewrite mtree_hashroot_ctree_hashroot. 451 | rewrite mtree_hashroot_ctree_hashroot. 452 | rewrite ctree_hashroot_ctree_H. 453 | f_equal. exact H2ad. 454 | - exact H2g. 455 | } 456 | destruct H2j as [T [H2ja H2jb]]. 457 | change (ledgerroot_valid (newledgerroot bh)). 458 | rewrite <- H2ja. 459 | unfold ledgerroot_valid. 460 | unfold ctree_valid. 461 | assert (L1: octree_supports_tx (tx_of_Block (bh,bd)) (Some (ctree_of_Block (bh,bd))) 0 (rewfn 0)). 462 | { destruct H2 as [H2x H2y]. 463 | exact (tx_of_Block_supported 0 (rewfn 0) check_hit targetf _ _ (bh,bd) LT H2x). 464 | } 465 | apply (mtree_hashroot_eq_valid (ctree_mtree T)). 466 | + rewrite mtree_hashroot_ctree_hashroot. 467 | rewrite mtree_hashroot_ctree_hashroot. 468 | rewrite ctree_hashroot_ctree_H. 469 | reflexivity. 470 | + destruct H1 as [f [H1a H1b]]. 471 | assert (LTf: octree_approx_fun_p (Some (ctree_of_Block (bh,bd))) f). 472 | { exact (endledgerroot_plr_valid_lem1 genlr bh bd f H1b H2ad H2g). } 473 | exists (tx_statefun_trans (tx_of_Block (bh,bd)) f). split. 474 | * { apply sf_tx_valid_thm with (fee := 0) (rew := (rewfn 0)). 475 | - exact H1a. 476 | - destruct H2 as [H2x H2y]. revert H2x. apply tx_of_Block_valid. 477 | - assert (L2: mtree_supports_tx (tx_of_Block (bh, bd)) 478 | (octree_mtree (Some (ctree_of_Block (bh, bd)))) 0 479 | (rewfn 0)). 480 | { revert L1. apply octree_mtree_supports_tx. } 481 | apply mtree_supports_tx_statefun with (f := f) in L2. 482 | + exact L2. 483 | + destruct H1a as [_ [Hf2 _]]. exact Hf2. 484 | + unfold octree_approx_fun_p in LTf. 485 | exact LTf. 486 | } 487 | * { assert (L4: mtree_approx_fun_p (octree_mtree (Some T)) (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 488 | { rewrite <- H2jb. 489 | set (tx' := (tx_of_Block (bh,bd))). 490 | set (T' := (ctree_of_Block (bh,bd))). 491 | assert (L5: octree_approx_fun_p (tx_octree_trans tx' (Some T')) (tx_statefun_trans tx' f)). 492 | { apply (octree_approx_trans tx' (Some T') f 0 (rewfn 0)). 493 | - exact H1a. 494 | - unfold tx'. unfold T'. destruct H2 as [H2x H2y]. revert H2x. apply tx_of_Block_supported. 495 | exact LT. 496 | - unfold T'. exact LTf. 497 | } 498 | unfold octree_approx_fun_p in L5. 499 | exact L5. 500 | } 501 | unfold octree_mtree in L4. 502 | exact L4. 503 | } 504 | - intros [ti' [H3 [H2 Hti]]]. 505 | generalize H2. 506 | intros [[H2aa [H2ab [H2ac [H2ad H2ae]]]] [H2af [H2ag [H2b [H2c [H2d [H2e [H2f [H2g [H2h [H2i H2j]]]]]]]]]]]. 507 | assert (LT: ctree_valid (ctree_of_Block (bh,bd))). 508 | { unfold ctree_of_Block. apply ctree_valid_cgraft_valid. 509 | - generalize (IH ti' H3). unfold ledgerroot_valid. unfold ctree_valid. 510 | apply mtree_hashroot_eq_valid. 511 | rewrite mtree_hashroot_ctree_hashroot. 512 | rewrite mtree_hashroot_ctree_hashroot. 513 | rewrite ctree_hashroot_ctree_H. 514 | f_equal. exact H2ad. 515 | - exact H2g. 516 | } 517 | destruct H2j as [T [H2ja H2jb]]. 518 | change (ledgerroot_valid (newledgerroot bh)). 519 | rewrite <- H2ja. 520 | unfold ledgerroot_valid. 521 | unfold ctree_valid. 522 | assert (L1: octree_supports_tx (tx_of_Block (bh,bd)) (Some (ctree_of_Block (bh,bd))) 0 (rewfn (S n))). 523 | { exact (tx_of_Block_supported (S n) (rewfn (S n)) check_hit targetf _ _ (bh,bd) LT H2). } 524 | apply (mtree_hashroot_eq_valid (ctree_mtree T)). 525 | + rewrite mtree_hashroot_ctree_hashroot. 526 | rewrite mtree_hashroot_ctree_hashroot. 527 | rewrite ctree_hashroot_ctree_H. 528 | reflexivity. 529 | + destruct (IH ti' H3) as [f [H1a H1b]]. 530 | assert (LTf: octree_approx_fun_p (Some (ctree_of_Block (bh,bd))) f). 531 | { exact (endledgerroot_plr_valid_lem1 _ bh bd f H1b H2ad H2g). } 532 | exists (tx_statefun_trans (tx_of_Block (bh,bd)) f). split. 533 | * { apply sf_tx_valid_thm with (fee := 0) (rew := (rewfn (S n))). 534 | - exact H1a. 535 | - revert H2. apply tx_of_Block_valid. 536 | - assert (L2: mtree_supports_tx (tx_of_Block (bh, bd)) 537 | (octree_mtree (Some (ctree_of_Block (bh, bd)))) 0 538 | (rewfn (S n))). 539 | { revert L1. apply octree_mtree_supports_tx. } 540 | apply mtree_supports_tx_statefun with (f := f) in L2. 541 | + exact L2. 542 | + destruct H1a as [_ [Hf2 _]]. exact Hf2. 543 | + unfold octree_approx_fun_p in LTf. 544 | exact LTf. 545 | } 546 | * { assert (L4: mtree_approx_fun_p (octree_mtree (Some T)) (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 547 | { rewrite <- H2jb. 548 | set (tx' := (tx_of_Block (bh,bd))). 549 | set (T' := (ctree_of_Block (bh,bd))). 550 | assert (L5: octree_approx_fun_p (tx_octree_trans tx' (Some T')) (tx_statefun_trans tx' f)). 551 | { apply (octree_approx_trans tx' (Some T') f 0 (rewfn (S n))). 552 | - exact H1a. 553 | - unfold tx'. unfold T'. revert H2. apply tx_of_Block_supported. 554 | exact LT. 555 | - unfold T'. exact LTf. 556 | } 557 | unfold octree_approx_fun_p in L5. 558 | exact L5. 559 | } 560 | unfold octree_mtree in L4. 561 | exact L4. 562 | } 563 | Qed. 564 | 565 | Fixpoint units_as_of_block (initdistr:nat) (rewfn:nat -> nat) (blockheight:nat) : nat := 566 | match blockheight with 567 | | O => initdistr 568 | | S n => units_as_of_block initdistr rewfn n + rewfn n 569 | end. 570 | 571 | Opaque sf_valid. 572 | Opaque tx_statefun_trans. 573 | Opaque tx_octree_trans. 574 | Opaque mtree_approx_fun_p. 575 | 576 | Theorem lastledgerroot_sum_correct initdistr al bl rewfn check_hit targetf genlr genti ti {n} 577 | (bc : BlockChain n) : 578 | ledgerroot_valid genlr -> 579 | octree_totalassets (Some (ctreeH 160 genlr)) al -> 580 | asset_value_sum al = initdistr -> 581 | valid_BlockChain rewfn check_hit targetf genlr genti ti bc -> 582 | octree_totalassets (Some (ctreeH 160 (lastledgerroot_of_BlockChain bc))) bl -> 583 | asset_value_sum bl = units_as_of_block initdistr rewfn (S n). 584 | intros H1 H2 H3. revert bl ti. 585 | induction bc as [[bh bd]|n bc IH [bh bd]]. 586 | - intros bl ti H4 H5. 587 | simpl. unfold valid_BlockChain in H4. 588 | generalize H4. 589 | intros [[[H4aa [H4ab [H4ac [H4ad H4ae]]]] [H4af [H4ag [H4b [H4c [H4d [H4e [H4f [H4g [H4h [H4i [T [HT1 HT2]]]]]]]]]]]]] Hti]. 590 | assert (L0: ctree_hashroot (ctree_of_Block(bh,bd)) = genlr). 591 | { transitivity (ctree_hashroot (prevledger bh)). 592 | - unfold ctree_of_Block. symmetry. 593 | apply subqc_hashroot_eq. apply ctree_cgraft_subqc. exact H4g. 594 | - exact H4ad. 595 | } 596 | assert (L1: octree_valid (Some (ctree_of_Block (bh, bd)))). 597 | { unfold octree_valid. revert H1. unfold ledgerroot_valid. unfold ctree_valid. 598 | apply mtree_hashroot_eq_valid. 599 | rewrite mtree_hashroot_ctree_hashroot. 600 | rewrite mtree_hashroot_ctree_hashroot. 601 | rewrite L0. 602 | rewrite ctree_hashroot_ctree_H. 603 | reflexivity. 604 | } 605 | assert (L2: tx_valid (tx_of_Block (bh, bd))). 606 | { destruct H4 as [H4x H4y]. revert H4x. apply tx_of_Block_valid. } 607 | assert (L3: octree_supports_tx (tx_of_Block (bh, bd)) (Some (ctree_of_Block (bh, bd))) 0 (rewfn 0)). 608 | { unfold octree_supports_tx. 609 | destruct H4 as [H4x H4y]. 610 | revert L1 H4x. unfold octree_valid. apply tx_of_Block_supported. 611 | } 612 | generalize L1. intros [f [Hf HTf]]. 613 | assert (L4: mtree_approx_fun_p (ctree_mtree (ctreeH 160 genlr)) f). 614 | { revert HTf. apply mtree_hashroot_mtree_approx_fun_p. 615 | rewrite mtree_hashroot_ctree_hashroot. rewrite L0. 616 | rewrite mtree_hashroot_ctree_hashroot. 617 | rewrite ctree_hashroot_ctree_H. 618 | reflexivity. 619 | } 620 | assert (L5: octree_totalassets (Some (ctree_of_Block (bh, bd))) al). 621 | { unfold octree_totalassets. unfold octree_mtree. 622 | apply (mtree_approx_fun_p_totalassets _ _ _ HTf). 623 | revert H2. unfold octree_totalassets. unfold octree_mtree. 624 | revert L4. clear. apply mtree_approx_fun_p_totalassets. 625 | } 626 | assert (L6: octree_approx_fun_p 627 | (tx_octree_trans (tx_of_Block (bh, bd)) 628 | (Some (ctree_of_Block (bh, bd)))) 629 | (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 630 | { apply (octree_approx_trans (tx_of_Block (bh,bd)) (Some (ctree_of_Block (bh,bd))) f 0 (rewfn 0) Hf L3). 631 | unfold octree_approx_fun_p. unfold octree_mtree. 632 | exact HTf. 633 | } 634 | assert (L7: octree_approx_fun_p (Some T) (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 635 | { unfold octree_approx_fun_p. rewrite <- HT2. 636 | exact L6. 637 | } 638 | assert (L8: octree_totalassets (tx_octree_trans (tx_of_Block (bh, bd)) (Some (ctree_of_Block (bh, bd)))) bl). 639 | { rewrite HT2. 640 | unfold octree_totalassets. unfold octree_mtree. 641 | apply (mtree_approx_fun_p_totalassets _ _ _ L7). 642 | revert H5. 643 | unfold octree_totalassets. unfold octree_mtree. 644 | apply mtree_approx_fun_p_totalassets. 645 | revert L7. unfold octree_approx_fun_p. 646 | apply mtree_hashroot_mtree_approx_fun_p. 647 | unfold octree_mtree. 648 | rewrite mtree_hashroot_ctree_hashroot. 649 | rewrite mtree_hashroot_ctree_hashroot. 650 | rewrite ctree_hashroot_ctree_H. 651 | rewrite HT1. reflexivity. 652 | } 653 | generalize (octree_tx_asset_value_sum (Some (ctree_of_Block (bh,bd))) (tx_of_Block (bh,bd)) 0 (rewfn 0) al bl L1 L2 L3 L5 L8). 654 | rewrite H3. 655 | assert (L9: asset_value_sum bl + 0 = asset_value_sum bl). 656 | { clear. omega. } 657 | rewrite L9. exact (fun H => H). 658 | - intros bl ti [ti' [H4 [H5 Hti]]] H6. 659 | generalize H5. 660 | intros [[H5aa [H5ab [H5ac [H5ad H5ae]]]] [H5af [H5ag [H5b [H5c [H5d [H5e [H5f [H5g [H5h [H5i [T [HT1 HT2]]]]]]]]]]]]]. 661 | assert (Lplr: ledgerroot_valid (lastledgerroot_of_BlockChain bc)). 662 | { revert H1 H4. apply lastledgerroot_valid. } 663 | generalize Lplr. intros [f [Hf Hplrf]]. 664 | assert (LIH: asset_value_sum (totalassets_ f) = units_as_of_block initdistr rewfn (S n)). 665 | { apply (IH (totalassets_ f) ti'). 666 | - exact H4. 667 | - unfold octree_totalassets. unfold octree_mtree. 668 | apply (mtree_approx_fun_p_totalassets _ _ _ Hplrf). 669 | reflexivity. 670 | } 671 | change (asset_value_sum bl = units_as_of_block initdistr rewfn (S n) + rewfn (S n)). 672 | rewrite <- LIH. 673 | assert (L0: ctree_hashroot (ctree_of_Block(bh,bd)) = lastledgerroot_of_BlockChain bc). 674 | { transitivity (ctree_hashroot (prevledger bh)). 675 | - unfold ctree_of_Block. symmetry. 676 | apply subqc_hashroot_eq. apply ctree_cgraft_subqc. exact H5g. 677 | - exact H5ad. 678 | } 679 | assert (L1: octree_valid (Some (ctree_of_Block (bh, bd)))). 680 | { unfold octree_valid. revert Lplr. unfold ledgerroot_valid. unfold ctree_valid. 681 | apply mtree_hashroot_eq_valid. 682 | rewrite mtree_hashroot_ctree_hashroot. 683 | rewrite mtree_hashroot_ctree_hashroot. 684 | rewrite L0. 685 | rewrite ctree_hashroot_ctree_H. 686 | reflexivity. 687 | } 688 | assert (L2: tx_valid (tx_of_Block (bh, bd))). 689 | { revert H5. apply tx_of_Block_valid. } 690 | assert (L3: octree_supports_tx (tx_of_Block (bh, bd)) (Some (ctree_of_Block (bh, bd))) 0 (rewfn (S n))). 691 | { unfold octree_supports_tx. revert L1 H5. unfold octree_valid. apply tx_of_Block_supported. } 692 | assert (L4: mtree_approx_fun_p (ctree_mtree (ctree_of_Block (bh,bd))) f). 693 | { revert Hplrf. apply mtree_hashroot_mtree_approx_fun_p. 694 | rewrite mtree_hashroot_ctree_hashroot. 695 | rewrite mtree_hashroot_ctree_hashroot. rewrite L0. 696 | rewrite ctree_hashroot_ctree_H. 697 | reflexivity. 698 | } 699 | assert (L5: octree_totalassets (Some (ctree_of_Block (bh, bd))) (totalassets_ f)). 700 | { unfold octree_totalassets. unfold octree_mtree. 701 | apply (mtree_approx_fun_p_totalassets _ _ _ L4). 702 | reflexivity. 703 | } 704 | assert (L6: octree_approx_fun_p 705 | (tx_octree_trans (tx_of_Block (bh, bd)) 706 | (Some (ctree_of_Block (bh, bd)))) 707 | (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 708 | { apply (octree_approx_trans (tx_of_Block (bh,bd)) (Some (ctree_of_Block (bh,bd))) f 0 (rewfn (S n)) Hf L3). 709 | unfold octree_approx_fun_p. unfold octree_mtree. 710 | exact L4. 711 | } 712 | assert (L7: octree_approx_fun_p (Some T) (tx_statefun_trans (tx_of_Block (bh, bd)) f)). 713 | { unfold octree_approx_fun_p. rewrite <- HT2. 714 | exact L6. 715 | } 716 | assert (L8: octree_totalassets (tx_octree_trans (tx_of_Block (bh, bd)) (Some (ctree_of_Block (bh, bd)))) bl). 717 | { rewrite HT2. 718 | unfold octree_totalassets. unfold octree_mtree. 719 | apply (mtree_approx_fun_p_totalassets _ _ _ L7). 720 | revert H6. 721 | unfold octree_totalassets. unfold octree_mtree. 722 | apply mtree_approx_fun_p_totalassets. 723 | revert L7. unfold octree_approx_fun_p. 724 | apply mtree_hashroot_mtree_approx_fun_p. 725 | unfold octree_mtree. 726 | rewrite mtree_hashroot_ctree_hashroot. 727 | rewrite mtree_hashroot_ctree_hashroot. 728 | rewrite ctree_hashroot_ctree_H. 729 | rewrite HT1. reflexivity. 730 | } 731 | generalize (octree_tx_asset_value_sum (Some (ctree_of_Block (bh,bd))) (tx_of_Block (bh,bd)) 0 (rewfn (S n)) (totalassets_ f) bl L1 L2 L3 L5 L8). 732 | assert (L9: asset_value_sum bl + 0 = asset_value_sum bl). 733 | { clear. omega. } 734 | rewrite L9. exact (fun H => H). 735 | Qed. 736 | -------------------------------------------------------------------------------- /CTreeGrafting.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** CTreeGrafting: A cgraft describes how to extend a ctree to a better approximation. 6 | The purpose is so that a small ctree can be given in the block header and 7 | the graft to extend it to a larger ctree can be given in the block delta without 8 | repeating the ctree information in the header. 9 | **) 10 | 11 | Require Export CTrees. 12 | 13 | (*** 14 | A "cgraft" can be used to extend a ctree to be a better approximation without repeating the part that is known. 15 | The idea is to use this to put just what's needed in the block header 16 | (enough of an approx to support the coinstake tx) 17 | and then a graft of the extra part that's needed in the block delta. 18 | ***) 19 | Definition cgraft : Type := list (prod hashval (sigT ctree)). 20 | 21 | Fixpoint cgraft_valid (G:cgraft) : Prop := 22 | match G with 23 | | nil => True 24 | | (h,existT n T)::G' => ctree_hashroot T = h /\ cgraft_valid G' 25 | end. 26 | 27 | (*** Coercion needed due to dependent types ***) 28 | Definition eqnat_coerce_ctree (m n:nat) (E:m = n) (T:ctree m) : ctree n := 29 | eq_rect m ctree T n E. 30 | 31 | (*** 32 | The dependent types here are a little tricky. 33 | I could omit m from the input return (existT n T) where T, leaving the caller of the function to check if m = n (which it should). 34 | Instead I decided to do it here. Either way a coercion using the proof of the equation seems to be needed. 35 | ***) 36 | Fixpoint cgraft_assoc (G:cgraft) (k:hashval) (m:nat) : ctree m := 37 | match G with 38 | | nil => ctreeH m k 39 | | (h,existT n T)::G' => 40 | if hashval_eq_dec h k then 41 | match eq_nat_dec n m with 42 | | left E => eqnat_coerce_ctree n m E T 43 | | _ => ctreeH m k 44 | end 45 | else 46 | cgraft_assoc G' k m 47 | end. 48 | 49 | Lemma cgraft_assoc_hashroot (G:cgraft) (k:hashval) (m:nat) : 50 | cgraft_valid G -> 51 | ctree_hashroot (cgraft_assoc G k m) = k. 52 | induction G as [|[h [n T]] G' IH]. 53 | - intros _. destruct m; reflexivity. 54 | - simpl. intros [H1 H2]. destruct (hashval_eq_dec h k) as [E1|E1]. 55 | + destruct (eq_nat_dec n m) as [E2|E2]. 56 | * destruct E2. simpl. (*** tricky ***) 57 | rewrite <- E1. exact H1. 58 | * destruct m; reflexivity. (*** even though this is impossible (since if ctree m and ctree n have the same hashroot, m should equal n) the proof goes through. ***) 59 | + exact (IH H2). 60 | Qed. 61 | 62 | Lemma cgraft_assoc_subqc (G:cgraft) (k:hashval) (m:nat) : 63 | cgraft_valid G -> subqc (ctreeH m k) (cgraft_assoc G k m). 64 | intros H1. generalize (cgraft_assoc_hashroot G k m H1). intros H2. 65 | destruct m as [|m]. 66 | - simpl. unfold subqc. unfold subqm. apply subqhH. 67 | change (mtree_hashroot (ctree_mtree (cgraft_assoc G k 0)) = Some k). 68 | rewrite mtree_hashroot_ctree_hashroot. congruence. 69 | - unfold subqc. 70 | change (subqm (mtreeH m (Some k)) (ctree_mtree (cgraft_assoc G k (S m)))). 71 | change (mtree_hashroot (ctree_mtree (cgraft_assoc G k (S m))) = mtree_hashroot (mtreeH m (Some k))). 72 | change (mtree_hashroot (ctree_mtree (cgraft_assoc G k (S m))) = Some k). 73 | rewrite mtree_hashroot_ctree_hashroot. congruence. 74 | Qed. 75 | 76 | Fixpoint ctree_cgraft (G:cgraft) {n} : ctree n -> ctree n := 77 | match n with 78 | | O => fun T:ctree 0 => 79 | match T with 80 | | inl h => cgraft_assoc G h 0 81 | | _ => T 82 | end 83 | | S n => fun T:ctree (S n) => 84 | match T with 85 | | inr (inl h) => cgraft_assoc G h (S n) 86 | | inr (inr (inl Tl)) => ctreeBL (ctree_cgraft G Tl) 87 | | inr (inr (inr (inl Tr))) => ctreeBR (ctree_cgraft G Tr) 88 | | inr (inr (inr (inr (Tl,Tr)))) => ctreeB (ctree_cgraft G Tl) (ctree_cgraft G Tr) 89 | | _ => T 90 | end 91 | end. 92 | 93 | Theorem ctree_cgraft_subqc (G:cgraft) {n} (T:ctree n) : 94 | cgraft_valid G -> subqc T (ctree_cgraft G T). 95 | intros H1. induction n as [|n IH]. 96 | - destruct T as [h|[a hl]]. 97 | + simpl. 98 | change (subqc (ctreeH 0 h) (cgraft_assoc G h 0)). 99 | now apply cgraft_assoc_subqc. 100 | + apply subqm_ref. 101 | - destruct T as [[gamma hl]|[h|[Tl|[Tr|[Tl Tr]]]]]. 102 | + apply subqm_ref. 103 | + change (subqc (ctreeH (S n) h) (cgraft_assoc G h (S n))). 104 | now apply cgraft_assoc_subqc. 105 | + split. 106 | * apply IH. 107 | * apply subqm_ref. 108 | + split. 109 | * apply subqm_ref. 110 | * apply IH. 111 | + split. 112 | * apply IH. 113 | * apply IH. 114 | Qed. 115 | 116 | Transparent ctree_mtree. 117 | 118 | Theorem ctree_approx_fun_cgraft_valid (G:cgraft) {n} (T:ctree n) (f:bitseq n -> list asset) : 119 | cgraft_valid G -> 120 | octree_approx_fun_p (Some T) f -> 121 | octree_approx_fun_p (Some (ctree_cgraft G T)) f. 122 | intros HG. 123 | induction n as [|n IH]. 124 | - destruct T as [h|[h hl]]. 125 | + unfold octree_approx_fun_p. simpl. intros H1. 126 | rewrite H1. 127 | generalize (cgraft_assoc_hashroot G h 0 HG). 128 | destruct (cgraft_assoc G h 0) as [k|[k kl]]. 129 | * simpl. congruence. 130 | * simpl. destruct (hlist_hashroot kl); congruence. 131 | + simpl. tauto. 132 | - destruct T as [[gamma hl]|[h|[Tl|[Tr|[Tl Tr]]]]]. 133 | + simpl. tauto. 134 | + unfold octree_approx_fun_p. 135 | intros [Tl [Tr [H1 [H2 H3]]]]. 136 | unfold ctree_cgraft. 137 | generalize (cgraft_assoc_hashroot G h (S n) HG). 138 | intros H4. 139 | destruct (cgraft_assoc G h (S n)) as [[[[|] delta] kl]|[k|[Tl'|[Tr'|[Tl' Tr']]]]]. 140 | * { simpl. simpl in H4. 141 | destruct (mtree_hashroot Tl) as [hl|] eqn:E1; 142 | destruct (mtree_hashroot Tr) as [hr|] eqn:E2; 143 | simpl in H1. 144 | - exfalso. inversion H1. rewrite H0 in H4. 145 | apply hashpairinj in H4. destruct H4 as [H4 _]. 146 | apply hashnatinj in H4. omega. 147 | - exfalso. inversion H1. rewrite H0 in H4. 148 | apply hashpairinj in H4. destruct H4 as [H4 _]. 149 | apply hashnatinj in H4. omega. 150 | - inversion H1. rewrite H0 in H4. 151 | apply hashpairinj in H4. destruct H4 as [_ H4]. 152 | split. 153 | + apply (mtree_hashroot_mtree_approx_fun_p Tl). 154 | * rewrite mtree_hashroot_empty. exact E1. 155 | * exact H2. 156 | + apply (mtree_hashroot_mtree_approx_fun_p Tr). 157 | * assert (L1: mtree_hashroot Tr = mtree_hashroot (ctree_mtree (ctreeL kl delta))). 158 | { rewrite mtree_hashroot_ctree_hashroot. 159 | congruence. 160 | } 161 | destruct n; exact L1. 162 | * exact H3. 163 | - discriminate H1. 164 | } 165 | * { simpl. simpl in H4. 166 | destruct (mtree_hashroot Tl) as [hl|] eqn:E1; 167 | destruct (mtree_hashroot Tr) as [hr|] eqn:E2; 168 | simpl in H1. 169 | - exfalso. inversion H1. rewrite H0 in H4. 170 | apply hashpairinj in H4. destruct H4 as [H4 _]. 171 | apply hashnatinj in H4. omega. 172 | - inversion H1. rewrite H0 in H4. 173 | apply hashpairinj in H4. destruct H4 as [_ H4]. 174 | split. 175 | + apply (mtree_hashroot_mtree_approx_fun_p Tl). 176 | * assert (L1: mtree_hashroot Tl = mtree_hashroot (ctree_mtree (ctreeL kl delta))). 177 | { rewrite mtree_hashroot_ctree_hashroot. 178 | congruence. 179 | } 180 | destruct n; exact L1. 181 | * exact H2. 182 | + apply (mtree_hashroot_mtree_approx_fun_p Tr). 183 | * rewrite mtree_hashroot_empty. exact E2. 184 | * exact H3. 185 | - exfalso. inversion H1. rewrite H0 in H4. 186 | apply hashpairinj in H4. destruct H4 as [H4 _]. 187 | apply hashnatinj in H4. omega. 188 | - discriminate H1. 189 | } 190 | * { simpl. exists Tl. exists Tr. simpl in H4. subst k. repeat split. 191 | - exact H1. 192 | - exact H2. 193 | - exact H3. 194 | } 195 | * { simpl. simpl in H4. 196 | destruct (mtree_hashroot Tl) as [hl|] eqn:E1; 197 | destruct (mtree_hashroot Tr) as [hr|] eqn:E2; 198 | simpl in H1. 199 | - exfalso. inversion H1. rewrite H0 in H4. 200 | apply hashpairinj in H4. destruct H4 as [H4 _]. 201 | apply hashnatinj in H4. omega. 202 | - inversion H1. rewrite H0 in H4. 203 | apply hashpairinj in H4. destruct H4 as [_ H4]. 204 | split. 205 | + apply (mtree_hashroot_mtree_approx_fun_p Tl). 206 | * rewrite mtree_hashroot_ctree_hashroot. congruence. 207 | * exact H2. 208 | + apply (mtree_hashroot_mtree_approx_fun_p Tr). 209 | * rewrite mtree_hashroot_empty. exact E2. 210 | * exact H3. 211 | - exfalso. inversion H1. rewrite H0 in H4. 212 | apply hashpairinj in H4. destruct H4 as [H4 _]. 213 | apply hashnatinj in H4. omega. 214 | - exfalso. discriminate H1. 215 | } 216 | * { simpl. simpl in H4. 217 | destruct (mtree_hashroot Tl) as [hl|] eqn:E1; 218 | destruct (mtree_hashroot Tr) as [hr|] eqn:E2; 219 | simpl in H1. 220 | - exfalso. inversion H1. rewrite H0 in H4. 221 | apply hashpairinj in H4. destruct H4 as [H4 _]. 222 | apply hashnatinj in H4. omega. 223 | - exfalso. inversion H1. rewrite H0 in H4. 224 | apply hashpairinj in H4. destruct H4 as [H4 _]. 225 | apply hashnatinj in H4. omega. 226 | - inversion H1. rewrite H0 in H4. 227 | apply hashpairinj in H4. destruct H4 as [_ H4]. 228 | split. 229 | + apply (mtree_hashroot_mtree_approx_fun_p Tl). 230 | * rewrite mtree_hashroot_empty. exact E1. 231 | * exact H2. 232 | + apply (mtree_hashroot_mtree_approx_fun_p Tr). 233 | * rewrite mtree_hashroot_ctree_hashroot. congruence. 234 | * exact H3. 235 | - discriminate H1. 236 | } 237 | * { simpl. simpl in H4. 238 | destruct (mtree_hashroot Tl) as [hl|] eqn:E1; 239 | destruct (mtree_hashroot Tr) as [hr|] eqn:E2; 240 | simpl in H1. 241 | - inversion H1. rewrite H0 in H4. 242 | apply hashpairinj in H4. destruct H4 as [_ H4]. 243 | apply hashpairinj in H4. destruct H4 as [H4a H4]. 244 | apply hashpairinj in H4. destruct H4 as [H4b _]. 245 | split. 246 | + apply (mtree_hashroot_mtree_approx_fun_p Tl). 247 | * rewrite mtree_hashroot_ctree_hashroot. congruence. 248 | * exact H2. 249 | + apply (mtree_hashroot_mtree_approx_fun_p Tr). 250 | * rewrite mtree_hashroot_ctree_hashroot. congruence. 251 | * exact H3. 252 | - exfalso. inversion H1. rewrite H0 in H4. 253 | apply hashpairinj in H4. destruct H4 as [H4 _]. 254 | apply hashnatinj in H4. omega. 255 | - exfalso. inversion H1. rewrite H0 in H4. 256 | apply hashpairinj in H4. destruct H4 as [H4 _]. 257 | apply hashnatinj in H4. omega. 258 | - discriminate H1. 259 | } 260 | + intros [H1 H2]. split. 261 | * apply IH. exact H1. 262 | * exact H2. 263 | + intros [H1 H2]. split. 264 | * exact H1. 265 | * apply IH. exact H2. 266 | + intros [H1 H2]. split. 267 | * apply IH. exact H1. 268 | * apply IH. exact H2. 269 | Qed. 270 | 271 | Opaque ctree_cgraft. 272 | Opaque cgraft_valid. 273 | Opaque ctree_mtree. 274 | Opaque mtree_approx_fun_p. 275 | 276 | Theorem ctree_valid_cgraft_valid (G:cgraft) (T:ctree 160) : 277 | ctree_valid T -> 278 | cgraft_valid G -> 279 | ctree_valid (ctree_cgraft G T). 280 | intros [f [H1 H2]] HG. 281 | exists f. split. 282 | - exact H1. 283 | - assert (L1: octree_approx_fun_p (Some (ctree_cgraft G T)) f). 284 | { apply ctree_approx_fun_cgraft_valid. 285 | - exact HG. 286 | - unfold octree_approx_fun_p. unfold octree_mtree. exact H2. 287 | } 288 | unfold octree_approx_fun_p in L1. unfold octree_mtree in L1. 289 | exact L1. 290 | Qed. 291 | -------------------------------------------------------------------------------- /CryptoHashes.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** CryptoHashes : Idealized representation of cryptographic hashing 6 | using inductive types enclosed in a module. The module 7 | simulates the 'trapdoor' property by disallowing writing a destructor 8 | function via a match. **) 9 | 10 | Require Export Addrs. 11 | 12 | Module Type HashValsType. 13 | 14 | Parameter hashval : Type. 15 | Parameter hashnat : nat -> hashval. 16 | Parameter hashaddr : addr -> hashval. 17 | Parameter hashpair : hashval -> hashval -> hashval. 18 | 19 | Parameter hashval_eq_dec : forall (h1 h2:hashval), { h1 = h2 } + { h1 <> h2 }. 20 | 21 | Axiom hashnatinj : forall m n, hashnat m = hashnat n -> m = n. 22 | Axiom hashaddrinj : forall alpha beta, hashaddr alpha = hashaddr beta -> alpha = beta. 23 | Axiom hashpairinj : forall h1a h1b h2a h2b, hashpair h1a h1b = hashpair h2a h2b -> h1a = h2a /\ h1b = h2b. 24 | 25 | Axiom hashnataddrdiscr : forall m alpha, hashnat m <> hashaddr alpha. 26 | Axiom hashnatpairdiscr : forall m h1 h2, hashnat m <> hashpair h1 h2. 27 | Axiom hashaddrpairdiscr : forall alpha h1 h2, hashaddr alpha <> hashpair h1 h2. 28 | 29 | Axiom hashval_ind : forall p:hashval -> Prop, 30 | (forall n, p (hashnat n)) -> 31 | (forall alpha, p (hashaddr alpha)) -> 32 | (forall h1, p h1 -> forall h2, p h2 -> p (hashpair h1 h2)) -> 33 | forall h, p h. 34 | 35 | End HashValsType. 36 | 37 | Module HashVals : HashValsType. 38 | 39 | Inductive hashval' : Type := 40 | | hashnat' : nat -> hashval' 41 | | hashaddr' : addr -> hashval' 42 | | hashpair' : hashval' -> hashval' -> hashval'. 43 | 44 | Definition hashval := hashval'. 45 | Definition hashnat := hashnat'. 46 | Definition hashaddr := hashaddr'. 47 | Definition hashpair := hashpair'. 48 | 49 | Fixpoint hashval_eq_dec (h1 h2:hashval) : { h1 = h2 } + { h1 <> h2 }. 50 | destruct h1 as [n1|alpha1|h1a h1b]; destruct h2 as [n2|alpha2|h2a h2b]; try (right; discriminate). 51 | - destruct (eq_nat_dec n1 n2). 52 | + left. congruence. 53 | + right. congruence. 54 | - destruct (addr_eq_dec alpha1 alpha2). 55 | + left. congruence. 56 | + right. congruence. 57 | - destruct (hashval_eq_dec h1a h2a). 58 | + destruct (hashval_eq_dec h1b h2b). 59 | * left. congruence. 60 | * right. congruence. 61 | + right. congruence. 62 | Defined. 63 | 64 | Lemma hashnatinj : forall m n, hashnat m = hashnat n -> m = n. 65 | intros m n H. inversion H. reflexivity. 66 | Qed. 67 | 68 | Lemma hashaddrinj : forall alpha beta, hashaddr alpha = hashaddr beta -> alpha = beta. 69 | intros alpha beta H. inversion H. reflexivity. 70 | Qed. 71 | 72 | Lemma hashpairinj : forall h1a h1b h2a h2b, hashpair h1a h1b = hashpair h2a h2b -> h1a = h2a /\ h1b = h2b. 73 | intros h1a h1b h2a h2b H1. inversion H1. tauto. 74 | Qed. 75 | 76 | Lemma hashnataddrdiscr : forall m alpha, hashnat m <> hashaddr alpha. 77 | discriminate. 78 | Qed. 79 | 80 | Lemma hashnatpairdiscr : forall m h1 h2, hashnat m <> hashpair h1 h2. 81 | discriminate. 82 | Qed. 83 | 84 | Lemma hashaddrpairdiscr : forall alpha h1 h2, hashaddr alpha <> hashpair h1 h2. 85 | discriminate. 86 | Qed. 87 | 88 | Lemma hashval_ind : forall p:hashval -> Prop, 89 | (forall n, p (hashnat n)) -> 90 | (forall alpha, p (hashaddr alpha)) -> 91 | (forall h1, p h1 -> forall h2, p h2 -> p (hashpair h1 h2)) -> 92 | forall h, p h. 93 | exact hashval'_ind. 94 | Qed. 95 | 96 | End HashVals. 97 | 98 | Export HashVals. 99 | 100 | Lemma hashpair_neq_L h1 h2 : hashpair h1 h2 <> h1. 101 | revert h1 h2. apply (hashval_ind (fun h1 => forall h2, hashpair h1 h2 <> h1)). 102 | - intros n h2 H. symmetry in H. revert H. apply hashnatpairdiscr. 103 | - intros alpha h2 H. symmetry in H. revert H. apply hashaddrpairdiscr. 104 | - intros h1a IHa h1b IHb h2 H. apply hashpairinj in H. destruct H as [H _]. 105 | revert H. apply IHa. 106 | Qed. 107 | 108 | Lemma hashpair_neq_R h1 h2 : hashpair h1 h2 <> h2. 109 | revert h2 h1. apply (hashval_ind (fun h2 => forall h1, hashpair h1 h2 <> h2)). 110 | - intros n h1 H. symmetry in H. revert H. apply hashnatpairdiscr. 111 | - intros alpha h1 H. symmetry in H. revert H. apply hashaddrpairdiscr. 112 | - intros h2a IHa h2b IHb h1 H. apply hashpairinj in H. destruct H as [_ H]. 113 | revert H. apply IHb. 114 | Qed. 115 | 116 | Fixpoint hashlist (hl:list hashval) : hashval := 117 | match hl with 118 | | h::hr => hashpair h (hashlist hr) 119 | | nil => hashnat 0 120 | end. 121 | 122 | Lemma hashlistinj : forall hl1 hl2, hashlist hl1 = hashlist hl2 -> hl1 = hl2. 123 | intros hl1. induction hl1 as [|h1 hr1 IH]; intros [|h2 hr2]. 124 | - reflexivity. 125 | - simpl. intros H. exfalso. revert H. apply hashnatpairdiscr. 126 | - simpl. intros H. exfalso. symmetry in H. revert H. apply hashnatpairdiscr. 127 | - simpl. intros H3. apply hashpairinj in H3. destruct H3 as [H4 H5]. 128 | subst h2. f_equal. 129 | apply IH. exact H5. 130 | Qed. 131 | 132 | Lemma hashmapinj {X} (f : X -> hashval) (l l': list X) : 133 | (forall x y, f x = f y -> x = y) -> 134 | map f l = map f l' -> l = l'. 135 | intros H1. revert l'. induction l as [|x l IH]; intros [|y l']. 136 | - tauto. 137 | - discriminate. 138 | - discriminate. 139 | - intros H2. inversion H2. f_equal. 140 | + revert H0. apply H1. 141 | + revert H3. apply IH. 142 | Qed. 143 | 144 | Definition hashopair (h1 h2:option hashval) : option hashval := 145 | match h1,h2 with 146 | | None,None => None 147 | | Some h1,None => Some (hashpair (hashnat 0) h1) 148 | | None,Some h2 => Some (hashpair (hashnat 1) h2) 149 | | Some h1,Some h2 => Some (hashlist (hashnat 2::h1::h2::nil)) 150 | end. 151 | 152 | Lemma hashopairinj h1a h1b h2a h2b : hashopair h1a h1b = hashopair h2a h2b -> h1a = h2a /\ h1b = h2b. 153 | destruct h1a as [h1a|]; destruct h2a as [h2a|]; 154 | destruct h1b as [h1b|]; destruct h2b as [h2b|]; simpl; intros H; try (discriminate H). 155 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashpairinj in H3. 156 | destruct H3 as [H4 H5]. apply hashpairinj in H5. destruct H5 as [H6 H7]. 157 | split; congruence. 158 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 159 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 160 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. 161 | split; congruence. 162 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 163 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 164 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 165 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. apply hashnatinj in H2. discriminate H2. 166 | - inversion H. apply hashpairinj in H1. destruct H1 as [H2 H3]. 167 | split; congruence. 168 | - split; reflexivity. 169 | Qed. 170 | 171 | Lemma hashopair_None_1 (h1 h2:option hashval) : hashopair h1 h2 = None -> h1 = None. 172 | destruct h1 as [h1|]. 173 | - simpl. destruct h2; discriminate. 174 | - tauto. 175 | Qed. 176 | 177 | Lemma hashopair_None_2 (h1 h2:option hashval) : hashopair h1 h2 = None -> h2 = None. 178 | destruct h2 as [h2|]. 179 | - destruct h1; discriminate. 180 | - tauto. 181 | Qed. 182 | 183 | Definition hashopair1 (h1:hashval) (h2:option hashval) : hashval := 184 | match h2 with 185 | | Some h2 => hashlist (hashnat 2::h1::h2::nil) 186 | | None => hashpair (hashnat 0) h1 187 | end. 188 | 189 | Definition hashopair2 (h1:option hashval) (h2:hashval) : hashval := 190 | match h1 with 191 | | Some h1 => hashlist (hashnat 2::h1::h2::nil) 192 | | None => hashpair (hashnat 1) h2 193 | end. 194 | 195 | Fixpoint ohashlist (hl:list hashval) : option hashval := 196 | match hl with 197 | | nil => None 198 | | h::hr => 199 | match ohashlist hr with 200 | | None => Some(hashpair (hashnat 3) h) 201 | | Some k => Some(hashpair (hashnat 4) (hashpair h k)) 202 | end 203 | end. 204 | 205 | Lemma ohashlistinj : forall hl1 hl2, ohashlist hl1 = ohashlist hl2 -> hl1 = hl2. 206 | intros hl1. induction hl1 as [|h1 hr1 IH]; intros [|h2 hr2] H1. 207 | - reflexivity. 208 | - exfalso. simpl in H1. destruct (ohashlist hr2); discriminate H1. 209 | - exfalso. simpl in H1. destruct (ohashlist hr1); discriminate H1. 210 | - simpl in H1. 211 | destruct (ohashlist hr1) as [k1|] eqn:E1; destruct (ohashlist hr2) as [k2|] eqn:E2. 212 | + inversion H1. 213 | apply hashpairinj in H0. destruct H0 as [_ H0]. 214 | apply hashpairinj in H0. destruct H0 as [H0 H2]. 215 | subst h2. f_equal. apply IH. congruence. 216 | + exfalso. inversion H1. 217 | apply hashpairinj in H0. destruct H0 as [H0 _]. 218 | apply hashnatinj in H0. omega. 219 | + exfalso. inversion H1. 220 | apply hashpairinj in H0. destruct H0 as [H0 _]. 221 | apply hashnatinj in H0. omega. 222 | + inversion H1. 223 | apply hashpairinj in H0. destruct H0 as [_ H0]. 224 | subst h2. f_equal. apply IH. congruence. 225 | Qed. 226 | 227 | Inductive subh (h:hashval) : hashval -> Prop := 228 | | subh_L h' : subh h (hashpair h h') 229 | | subh_R h' : subh h (hashpair h' h) 230 | | subh_PL h1 h2 : subh h h1 -> subh h (hashpair h1 h2) 231 | | subh_PR h1 h2 : subh h h2 -> subh h (hashpair h1 h2) 232 | . 233 | 234 | Lemma subh_hashlist h hl : In h hl -> subh h (hashlist hl). 235 | induction hl. 236 | - simpl; tauto. 237 | - intros [H1|H1]; simpl. 238 | + subst a. apply subh_L. 239 | + apply subh_PR. now apply IHhl. 240 | Qed. 241 | 242 | Lemma subh_tra h1 h2 h3 : subh h1 h2 -> subh h2 h3 -> subh h1 h3. 243 | intros H. revert h3. induction H as [h2|h2|h2 h3 H1 IH1|h2 h3 H1 IH1]. 244 | - intros h3 H. induction H as [h3|h3|h3 h4 H2 IH2|h3 h4 H2 IH2]. 245 | + apply subh_PL. apply subh_L. 246 | + apply subh_PR. apply subh_L. 247 | + now apply subh_PL. 248 | + now apply subh_PR. 249 | - intros h3 H. induction H as [h3|h3|h3 h4 H2 IH2|h3 h4 H2 IH2]. 250 | + apply subh_PL. apply subh_R. 251 | + apply subh_PR. apply subh_R. 252 | + now apply subh_PL. 253 | + now apply subh_PR. 254 | - intros h4 H. induction H as [h4|h4|h4 h5 H2 IH2|h4 h5 H2 IH2]. 255 | + apply IH1. apply subh_PL. apply subh_L. 256 | + apply IH1. apply subh_PR. apply subh_L. 257 | + now apply subh_PL. 258 | + now apply subh_PR. 259 | - intros h4 H. induction H as [h4|h4|h4 h5 H2 IH2|h4 h5 H2 IH2]. 260 | + apply IH1. apply subh_PL. apply subh_R. 261 | + apply IH1. apply subh_PR. apply subh_R. 262 | + now apply subh_PL. 263 | + now apply subh_PR. 264 | Qed. 265 | 266 | Lemma subh_irrefl h : ~ subh h h. 267 | revert h. apply hashval_ind. 268 | - intros n H. inversion H. 269 | + symmetry in H1. revert H1. apply hashnatpairdiscr. 270 | + symmetry in H1. revert H1. apply hashnatpairdiscr. 271 | + symmetry in H0. revert H0. apply hashnatpairdiscr. 272 | + symmetry in H0. revert H0. apply hashnatpairdiscr. 273 | - intros alpha H. inversion H. 274 | + symmetry in H1. revert H1. apply hashaddrpairdiscr. 275 | + symmetry in H1. revert H1. apply hashaddrpairdiscr. 276 | + symmetry in H0. revert H0. apply hashaddrpairdiscr. 277 | + symmetry in H0. revert H0. apply hashaddrpairdiscr. 278 | - intros h1 IH1 h2 IH2 H. inversion H. 279 | + revert H1. apply hashpair_neq_L. 280 | + revert H1. apply hashpair_neq_R. 281 | + apply IH1. apply subh_tra with (h2 := (hashpair h1 h2)). 282 | * apply subh_L. 283 | * apply hashpairinj in H0. destruct H0 as [H0a H0b]. congruence. 284 | + apply IH2. apply subh_tra with (h2 := (hashpair h1 h2)). 285 | * apply subh_R. 286 | * apply hashpairinj in H0. destruct H0 as [H0a H0b]. congruence. 287 | Qed. 288 | 289 | Lemma subh_asym h h' : subh h h' -> ~ subh h' h. 290 | intros H1 H2. apply (subh_irrefl h). 291 | now apply subh_tra with (h2 := h'). 292 | Qed. 293 | -------------------------------------------------------------------------------- /CryptoSignatures.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** CryptoSignatures: The trivial code here simulates cryptographic signatures 6 | using a module that exports a type of signatures and a relation check_signat. 7 | There are some additional functions and properties exported that are not used 8 | in the rest of the development. 9 | **) 10 | 11 | Require Export CryptoHashes. 12 | 13 | Module Type SignatType. 14 | 15 | Parameter signat : Type. 16 | Parameter privkey : Type. 17 | Parameter privkey_addr : privkey -> addr. 18 | Parameter sign : privkey -> hashval -> nat -> signat. 19 | Parameter check_signat : hashval -> signat -> addr -> Prop. 20 | 21 | Axiom privkey_addr_inj : forall key key', privkey_addr key = privkey_addr key' -> key = key'. 22 | Axiom signat_prop : forall r key alpha h, privkey_addr key = alpha -> check_signat h (sign key h r) alpha. 23 | 24 | End SignatType. 25 | 26 | Module Signat : SignatType. 27 | 28 | Definition signat := unit. 29 | Definition privkey := addr. 30 | Definition privkey_addr : privkey -> addr := fun alpha => alpha. 31 | Definition sign : privkey -> hashval -> nat -> signat := fun _ _ _ => tt. 32 | Definition check_signat : hashval -> signat -> addr -> Prop := fun _ _ _ => True. 33 | 34 | Theorem privkey_addr_inj (key key':privkey) : privkey_addr key = privkey_addr key' -> key = key'. 35 | intros H. exact H. 36 | Qed. 37 | 38 | Theorem signat_prop r key alpha h : privkey_addr key = alpha -> check_signat h (sign key h r) alpha. 39 | intros _. unfold check_signat. tauto. 40 | Qed. 41 | 42 | End Signat. 43 | 44 | Export Signat. 45 | 46 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | The MIT License (MIT) 2 | 3 | Ledger Theory Coq Development 4 | Copyright (c) 2015 Bill White (12pbhpqEg7cjaCLLcvdhJhBWGUQWkRK3zS) 5 | 6 | Permission is hereby granted, free of charge, to any person obtaining a copy 7 | of this software and associated documentation files (the 'Software'), to deal 8 | in the Software without restriction, including without limitation the rights 9 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 10 | copies of the Software, and to permit persons to whom the Software is 11 | furnished to do so, subject to the following conditions: 12 | 13 | The above copyright notice and this permission notice shall be included in 14 | all copies or substantial portions of the Software. 15 | 16 | THE SOFTWARE IS PROVIDED 'AS IS', WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 17 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 18 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 19 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 20 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 21 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN 22 | THE SOFTWARE. 23 | ---------------------- 24 | Signature of the part above the line (including the last newline). 25 | 26 | HLYbKD1ObQJwuEzIGGsPGY3sJd6vDmZ4m1UWW9yGEAZJ+Ny2stS/6dhUeM0tnn49wj8Ge5rAUpJsErdyAoY3/Wc= 27 | 28 | Bill White (12pbhpqEg7cjaCLLcvdhJhBWGUQWkRK3zS) 29 | -------------------------------------------------------------------------------- /LedgerStates.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** LedgerStates: States are represented here as ledger functions from addresses to lists of assets 6 | (statefun). A ledger function is valid (sf_valid and sf_valid_) if no address holds duplicate assets, 7 | assetids are associated with at most one asset, every asset corresponds to a transaction output, 8 | and every asset id for an asset that has been spent at some time in the past is not 9 | currently spendable. The function tx_statefun_trans transforms a ledger function to 10 | a new ledger function. If the current ledger function supports the transaction (statefun_supports_tx) 11 | and the current ledger function is valid, then the transformed ledger function is also valid (sf_tx_valid_thm). 12 | The total asset values in the ledger change precisely according to the fees and rewards 13 | in the transaction (totalunits_bdd). 14 | **) 15 | 16 | Require Export Transactions. 17 | 18 | Definition statefun : Type := addr -> list asset. 19 | 20 | Definition sf_unsp_txout {n} (f:bitseq n -> list asset) (h:hashval) : Prop := 21 | exists alpha, In_dom h (f alpha). 22 | 23 | (*** sf_spent f h means h is the hash of a txout that has been spent at some point by the state represented by f. ***) 24 | Inductive sf_spent {n} (f:bitseq n -> list asset) (h:hashval) : Prop := 25 | | sf_spent_1 inpl outpl i alpha : sf_unsp_txout f (hashpair (hashtx(inpl,outpl)) (hashnat i)) -> In (alpha,h) inpl -> sf_spent f h 26 | | sf_spent_R inpl outpl i alpha : sf_spent f (hashpair (hashtx(inpl,outpl)) (hashnat i)) -> In (alpha,h) inpl -> sf_spent f h 27 | . 28 | 29 | Definition sf_valid_ {n} (alphapre:bitseq n -> addr) (f:bitseq n -> list asset) := 30 | (*** No duplicate: List at f alpha has no duplicate entries. ***) 31 | (forall alpha, no_dups (f alpha)) 32 | /\ 33 | (*** Functionality: An unspent txout is only associated with one asset in one address. ***) 34 | (forall h alpha u alpha' u', In (h,u) (f alpha) -> In (h,u') (f alpha') -> alpha = alpha' /\ u = u') 35 | /\ 36 | (*** Existence: For every (h,u) in (f alpha), h comes from a txout with output of asset u to address alpha. ***) 37 | (forall h u alpha, In (h,u) (f alpha) 38 | -> exists inpl outpl i, 39 | h = hashpair (hashtx(inpl,outpl)) (hashnat i) 40 | /\ 41 | nth_error outpl i = value(alphapre alpha,u)) 42 | /\ 43 | (*** No Double Spending: Every txout that has been spent according to f is not an unspent txout of f. ***) 44 | (forall h, sf_spent f h -> ~ sf_unsp_txout f h) 45 | . 46 | 47 | Definition sf_valid (f:statefun) : Prop := sf_valid_ (fun alpha => alpha) f. 48 | 49 | Lemma sf_valid_fnl (f:statefun) : sf_valid f -> forall alpha, fnl (f alpha). 50 | intros [Hf1 [Hf2 _]] alpha. 51 | generalize (fun h u => Hf2 h alpha u alpha). 52 | generalize (Hf1 alpha). 53 | generalize (f alpha) as al. 54 | intros al H. induction H as [|[h u] hr H1 H2 IH]. 55 | - intros _. apply fnl_N. 56 | - intros H3. apply fnl_C. 57 | + intros H4. apply In_In_dom_lem in H4. destruct H4 as [v H5]. 58 | assert (L1: In (h,u) ((h,u)::hr)) by now left. 59 | assert (L2: In (h,v) ((h,u)::hr)) by now right. 60 | destruct (H3 _ _ _ L1 L2) as [_ H6]. subst v. contradiction (H1 H5). 61 | + apply IH. intros k v v' H4 H5. apply (H3 k). 62 | * now right. 63 | * now right. 64 | Qed. 65 | 66 | Inductive statefun_asset_value_in f : list addr_assetid -> nat -> Prop := 67 | | statefun_asset_value_in_nil : statefun_asset_value_in f nil 0 68 | | statefun_asset_value_in_cons h a u inpl alpha v : 69 | statefun_asset_value_in f inpl v -> 70 | In a (f alpha) -> 71 | asset_value a = Some u -> 72 | h = assetid a -> 73 | statefun_asset_value_in f ((alpha,h)::inpl) (u+v) 74 | | statefun_asset_value_in_skip h a inpl alpha v : 75 | statefun_asset_value_in f inpl v -> 76 | In a (f alpha) -> 77 | asset_value a = None -> 78 | h = assetid a -> 79 | statefun_asset_value_in f ((alpha,h)::inpl) v 80 | . 81 | 82 | Fixpoint totalassets_ {n} : forall (f:bitseq n -> list asset), list asset := 83 | match n with 84 | | O => fun (f:bitseq 0 -> list asset) => f tt 85 | | S n' => fun (f:bitseq (S n') -> list asset) => totalassets_ (fun bs => f (false,bs)) ++ totalassets_ (fun bs => f (true,bs)) 86 | end. 87 | 88 | Definition statefun_totalassets (f:statefun) : list asset := totalassets_ f. 89 | 90 | Definition statefun_totalunits (f:statefun) : nat := 91 | asset_value_sum (statefun_totalassets f). 92 | 93 | Definition statefun_supports_tx (f:addr -> list asset) (tx:Tx) (fee rew:nat) : Prop := 94 | (exists utot:nat, 95 | statefun_asset_value_in f (tx_inputs tx) utot 96 | /\ 97 | asset_value_out (tx_outputs tx) + fee = utot + rew) 98 | . 99 | 100 | Definition tx_statefun_trans (tx:Tx) (f:statefun) : statefun := 101 | fun alpha:addr => 102 | (new_assets alpha (tx_outputs tx) (hashtx tx) 0) 103 | ++ 104 | (remove_assets (f alpha) 105 | (get_spent alpha (tx_inputs tx))). 106 | 107 | Lemma tx_statefun_trans_lem_iff f inpl outpl alpha h obl u : 108 | In (h,(obl,u)) (tx_statefun_trans (inpl,outpl) f alpha) 109 | <-> 110 | In (h,(obl,u)) (f alpha) /\ ~In (alpha,h) inpl 111 | \/ 112 | exists i, nth_error outpl i = value (alpha,(obl,u)) /\ h = hashpair (hashtx(inpl,outpl)) (hashnat i). 113 | unfold tx_statefun_trans. split. 114 | - intros H1. apply in_app_iff in H1. destruct H1 as [H1|H1]. 115 | + apply new_assets_iff in H1. destruct H1 as [j [H2 H3]]. 116 | simpl in H2. right. exists j. split. 117 | * assumption. 118 | * exact H3. 119 | + apply remove_assets_iff in H1. destruct H1 as [H2 H3]. left. split. 120 | * assumption. 121 | * intros H4. apply H3. apply get_spent_iff. exact H4. 122 | - intros [[H1 H2]|[i [H1 H2]]]. 123 | + apply in_app_iff. right. apply remove_assets_iff. split. 124 | * assumption. 125 | * intros H3. apply H2. now apply get_spent_iff. 126 | + apply in_app_iff. left. apply new_assets_iff. exists i. simpl. split. 127 | * assumption. 128 | * exact H2. 129 | Qed. 130 | 131 | Lemma statefun_supports_tx_assets_In (f:statefun) (tx:Tx) fee rew alpha h : 132 | statefun_supports_tx f tx fee rew -> 133 | In (alpha,h) (tx_inputs tx) -> exists obl u, In (h,(obl,u)) (f alpha). 134 | intros [utot [H _]]. destruct tx as [inpl outpl]. 135 | simpl in *|-*. 136 | induction H as [|k [k' [obl' u']] u inpl beta v H1 IH H2 H3 H4|k [k' [obl' u']] inpl beta v H1 IH H2 H3 H4]. 137 | - simpl. tauto. 138 | - intros [H5|H5]. 139 | + inversion H5. simpl in H4. subst k'. subst k. exists obl'. exists u'. 140 | subst beta. exact H2. 141 | + tauto. 142 | - intros [H5|H5]. 143 | + inversion H5. simpl in H4. subst k'. subst k. exists obl'. exists u'. 144 | subst beta. exact H2. 145 | + tauto. 146 | Qed. 147 | 148 | (*** We need to know the tx has at least one input to ensure all its txouts are fresh. ***) 149 | Lemma sf_tx_valid_fresh_lem (f:statefun) (tx:Tx) fee rew : 150 | sf_valid f -> 151 | tx_inputs_valid (tx_inputs tx) -> 152 | statefun_supports_tx f tx fee rew -> 153 | forall i alpha, ~In_dom (hashpair (hashtx tx) (hashnat i)) (f alpha). 154 | destruct tx as [[|[beta hin] inpl] outpl]. 155 | - intros _ [Ht2 _]. exfalso. simpl in Ht2. congruence. 156 | - intros Hf Ht Hs i alpha H1. 157 | generalize Hf. intros [Hf1 [Hf2 [Hf3 Hf4]]]. 158 | apply (Hf4 hin). 159 | + apply (sf_spent_1 f hin ((beta,hin)::inpl) outpl i beta). 160 | * exists alpha. assumption. 161 | * now left. 162 | + exists beta. apply In_In_dom_lem. 163 | assert (L1: exists obl u, In (hin, (obl,u)) (f beta)). 164 | { apply (statefun_supports_tx_assets_In _ _ _ _ _ _ Hs). 165 | simpl. now left. 166 | } 167 | destruct L1 as [obl [u H2]]. 168 | exists (obl,u). exact H2. 169 | Qed. 170 | 171 | (*** We need to know the tx has at least one input to ensure all its txouts did not occur as previous spent txs. ***) 172 | Lemma sf_tx_valid_not_spent_lem (f:statefun) (tx:Tx) fee rew : 173 | sf_valid f -> 174 | tx_inputs_valid (tx_inputs tx) -> 175 | statefun_supports_tx f tx fee rew -> 176 | forall i, ~ sf_spent f (hashpair (hashtx tx) (hashnat i)). 177 | destruct tx as [[|[beta hin] inpl] outpl]. 178 | - intros _ [Ht2 _]. exfalso. simpl in Ht2. congruence. 179 | - intros Hf Ht Hs i H1. 180 | generalize Hf. intros [Hf1 [Hf2 [Hf3 Hf4]]]. 181 | apply (Hf4 hin). 182 | + apply (sf_spent_R f hin ((beta,hin)::inpl) outpl i beta). 183 | * exact H1. 184 | * now left. 185 | + exists beta. apply In_In_dom_lem. 186 | assert (L1: exists obl u, In (hin, (obl,u)) (f beta)). 187 | { apply (statefun_supports_tx_assets_In _ _ _ _ _ _ Hs). 188 | simpl. now left. 189 | } 190 | destruct L1 as [obl [u H2]]. 191 | exists (obl,u). exact H2. 192 | Qed. 193 | 194 | (** If a txout h was spent in the transformed state, then it was either already spent or was one of the inputs of the new tx. ***) 195 | Lemma sf_tx_valid_spent_lem inpl outpl f h fee rew : 196 | tx_inputs_valid inpl -> 197 | statefun_supports_tx f (inpl,outpl) fee rew -> 198 | sf_spent (tx_statefun_trans (inpl, outpl) f) h -> 199 | sf_spent f h \/ exists alpha, In (alpha,h) inpl. 200 | intros Ht Hs H. 201 | induction H as [h inpl' outpl' i alpha [beta H1] H2|h inpl' outpl' i alpha H1 IH1 H2]. 202 | - apply In_In_dom_lem in H1. destruct H1 as [u H1]. 203 | apply in_app_iff in H1. destruct H1 as [H1|H1]. 204 | + apply new_assets_iff in H1. destruct H1 as [j [H3 H4]]. 205 | right. exists alpha. 206 | assert (L1: In (hash_addr_assetid (alpha,h)) (map hash_addr_assetid inpl)). 207 | { apply hashpairinj in H4. destruct H4 as [H5 H6]. apply hashnatinj in H6. 208 | simpl in H5. apply hashpairinj in H5. 209 | destruct H5 as [H7 H8]. 210 | apply hashlistinj in H7. rewrite <- H7. 211 | apply in_map_iff. exists (alpha,h). tauto. } 212 | apply in_map_iff in L1. destruct L1 as [[alpha' h'] [L1a L1b]]. 213 | apply hash_addr_assetidinj in L1a. destruct L1a as [L1aa L1ab]. 214 | subst alpha' h'. assumption. 215 | + apply remove_assets_iff in H1. destruct H1 as [H3 H4]. 216 | left. apply (sf_spent_1 f h inpl' outpl' i alpha). 217 | * exists beta. apply In_In_dom_lem_2 in H3. assumption. 218 | * assumption. 219 | - left. destruct IH1 as [H3|[beta H3]]. 220 | + now apply (sf_spent_R f h inpl' outpl' i alpha). 221 | + apply (statefun_supports_tx_assets_In _ _ _ _ _ _ Hs) in H3. 222 | destruct H3 as [obl [u H3]]. 223 | apply (sf_spent_1 f h inpl' outpl' i alpha). 224 | * exists beta. apply In_In_dom_lem_2 in H3. assumption. 225 | * assumption. 226 | Qed. 227 | 228 | Theorem sf_tx_valid_thm (f:statefun) (tx:Tx) fee rew : 229 | sf_valid f -> 230 | tx_valid tx -> 231 | statefun_supports_tx f tx fee rew -> 232 | sf_valid (tx_statefun_trans tx f). 233 | intros Hf Ht Hs. 234 | destruct tx as [inpl outpl]. generalize Hf Ht. simpl. 235 | intros [Hf1 [Hf2 [Hf3 Hf4]]] Ht2. split. 236 | - intros alpha. simpl. unfold tx_statefun_trans. apply no_dups_app. 237 | + apply no_dups_new_assets. 238 | + apply no_dups_remove_assets. apply Hf1. 239 | + intros [h [obl u]] H1 H2. 240 | apply new_assets_iff in H1. apply remove_assets_iff in H2. 241 | destruct H1 as [j [H3 H4]]. destruct H2 as [H5 H6]. 242 | apply (sf_tx_valid_fresh_lem f (inpl,outpl) fee rew Hf Ht2 Hs j alpha). 243 | apply In_In_dom_lem. exists (obl,u). 244 | change (h = hashpair (hashtx (inpl, outpl)) (hashnat j)) in H4. 245 | rewrite H4 in H5. exact H5. 246 | - split. 247 | + intros h alpha [obl u] alpha' [obl' u']. 248 | unfold tx_statefun_trans. intros H1 H2. 249 | apply in_app_iff in H1. apply in_app_iff in H2. 250 | destruct H1 as [H1|H1]; destruct H2 as [H2|H2]. 251 | * apply new_assets_iff in H1. destruct H1 as [j [H3 H4]]. 252 | apply new_assets_iff in H2. destruct H2 as [k [H5 H6]]. 253 | change (nth_error outpl j = value (alpha, (obl,u))) in H3. 254 | change (h = hashpair (hashtx (inpl, outpl)) (hashnat j)) in H4. 255 | change (nth_error outpl k = value (alpha', (obl',u'))) in H5. 256 | change (h = hashpair (hashtx (inpl, outpl)) (hashnat k)) in H6. 257 | assert (L1: j = k). 258 | { subst h. apply hashpairinj in H6. destruct H6 as [H7 H8]. 259 | apply hashnatinj in H8. omega. } 260 | subst k. rewrite H3 in H5. inversion H5. tauto. 261 | * exfalso. 262 | apply new_assets_iff in H1. destruct H1 as [j [H3 H4]]. 263 | change (nth_error outpl j = value (alpha, (obl,u))) in H3. 264 | change (h = hashpair (hashtx (inpl, outpl)) (hashnat j)) in H4. 265 | apply remove_assets_iff in H2. destruct H2 as [H5 H6]. 266 | change (~ In h (get_spent alpha' inpl)) in H6. 267 | apply In_In_dom_lem_2 in H5. revert H5. 268 | rewrite H4. apply (sf_tx_valid_fresh_lem f (inpl,outpl) fee rew Hf Ht2 Hs). 269 | * exfalso. 270 | apply remove_assets_iff in H1. destruct H1 as [H3 H4]. 271 | change (~ In h (get_spent alpha inpl)) in H4. 272 | apply new_assets_iff in H2. destruct H2 as [j [H5 H6]]. 273 | change (nth_error outpl j = value (alpha', (obl',u'))) in H5. 274 | change (h = hashpair (hashtx (inpl, outpl)) (hashnat j)) in H6. 275 | apply In_In_dom_lem_2 in H3. revert H3. 276 | rewrite H6. apply (sf_tx_valid_fresh_lem f (inpl,outpl) fee rew Hf Ht2 Hs). 277 | * apply remove_assets_iff in H1. destruct H1 as [H3 H4]. 278 | apply remove_assets_iff in H2. destruct H2 as [H5 H6]. 279 | change (~ In h (get_spent alpha inpl)) in H4. 280 | change (~ In h (get_spent alpha' inpl)) in H6. 281 | now apply Hf2 with (h := h). 282 | + split. 283 | * { unfold tx_statefun_trans. intros h [obl u] alpha H1. 284 | apply in_app_iff in H1. destruct H1 as [H1|H1]. 285 | - apply new_assets_iff in H1. destruct H1 as [j [H2 H3]]. 286 | exists inpl. exists outpl. exists j. split. 287 | + assumption. 288 | + exact H2. 289 | - apply remove_assets_iff in H1. destruct H1 as [H2 H3]. 290 | revert H2. apply Hf3. 291 | } 292 | * { intros h H1 [alpha H2]. 293 | apply In_In_dom_lem in H2. destruct H2 as [u H2]. 294 | apply in_app_iff in H2. destruct H2 as [H2|H2]. 295 | - apply new_assets_iff in H2. destruct H2 as [j [H3 H4]]. 296 | simpl in H3, H4. 297 | destruct (sf_tx_valid_spent_lem inpl outpl f h fee rew Ht2 Hs H1) as [H5|[beta H5]]. 298 | + revert H5. rewrite H4. 299 | apply (sf_tx_valid_not_spent_lem f (inpl,outpl) fee rew Hf Ht2 Hs). 300 | + apply (hashtx_notin_inpl beta inpl outpl j). 301 | simpl. rewrite <- H4. exact H5. 302 | - apply remove_assets_iff in H2. destruct H2 as [H3 H4]. 303 | assert (L1: sf_unsp_txout f h). 304 | { exists alpha. apply In_In_dom_lem_2 in H3. assumption. } 305 | revert L1. apply Hf4. 306 | destruct (sf_tx_valid_spent_lem inpl outpl f h fee rew Ht2 Hs H1) as [H5|[beta H5]]. 307 | + assumption. 308 | + exfalso. apply H4. apply get_spent_iff. simpl. 309 | assert (L2: alpha = beta). 310 | { destruct (statefun_supports_tx_assets_In f (inpl,outpl) fee rew beta h Hs H5) as [obl' [v H6]]. 311 | generalize (Hf2 h alpha u beta (obl',v) H3 H6). tauto. 312 | } 313 | congruence. 314 | } 315 | Qed. 316 | 317 | Lemma totalassets__iff {n} (f:bitseq n -> list asset) h u : 318 | In (h,u) (totalassets_ f) 319 | <-> 320 | exists bs, In (h,u) (f bs). 321 | induction n as [|n IH]. 322 | - simpl. split. 323 | + intros H1. exists tt. assumption. 324 | + intros [[] H1]. assumption. 325 | - simpl. split. 326 | + intros H1. apply in_app_iff in H1. destruct H1 as [H1|H1]. 327 | * apply IH in H1. destruct H1 as [bs H1]. exists (false,bs). assumption. 328 | * apply IH in H1. destruct H1 as [bs H1]. exists (true,bs). assumption. 329 | + intros [[[|] bs] H1]. 330 | * apply in_or_app. right. apply IH. exists bs. assumption. 331 | * apply in_or_app. left. apply IH. exists bs. assumption. 332 | Qed. 333 | 334 | Lemma totalassets_iff (f:addr -> list asset) h u : 335 | In (h,u) (statefun_totalassets f) 336 | <-> 337 | exists alpha, In (h,u) (f alpha). 338 | apply totalassets__iff. 339 | Qed. 340 | 341 | Lemma totalassets_no_dups_ {n} (f:bitseq n -> list asset) : 342 | (forall bs, no_dups (f bs)) -> 343 | (forall h bs u bs' u', In (h,u) (f bs) -> In (h,u') (f bs') -> bs = bs' /\ u = u') -> 344 | no_dups (totalassets_ f). 345 | induction n as [|n IH]. 346 | - simpl. intros H1 _. apply H1. 347 | - intros H1 H2. simpl. apply no_dups_app. 348 | + apply IH. 349 | * intros bs. apply (H1 (false,bs)). 350 | * intros h bs u bs' u' H3 H4. 351 | destruct (H2 h (false,bs) u (false,bs') u' H3 H4) as [H5 H6]. 352 | split; congruence. 353 | + apply IH. 354 | * intros bs. apply (H1 (true,bs)). 355 | * intros h bs u bs' u' H3 H4. 356 | destruct (H2 h (true,bs) u (true,bs') u' H3 H4) as [H5 H6]. 357 | split; congruence. 358 | + intros [h u] H3 H4. 359 | apply totalassets__iff in H3. apply totalassets__iff in H4. 360 | destruct H3 as [bs H3]. destruct H4 as [bs' H4]. 361 | destruct (H2 h (false,bs) u (true,bs') u H3 H4) as [H5 H6]. 362 | discriminate H5. 363 | Qed. 364 | 365 | Lemma totalassets_no_dups (f:statefun) : 366 | sf_valid f -> 367 | no_dups (statefun_totalassets f). 368 | intros [Hf1 [Hf2 _]]. apply totalassets_no_dups_. 369 | - exact Hf1. 370 | - exact Hf2. 371 | Qed. 372 | 373 | Lemma totalassets_fnl_ {n} (f:bitseq n -> list asset) : 374 | (forall bs, no_dups (f bs)) -> 375 | (forall h bs u bs' u', In (h,u) (f bs) -> In (h,u') (f bs') -> bs = bs' /\ u = u') -> 376 | fnl (totalassets_ f). 377 | induction n as [|n IH]. 378 | - simpl. intros H1 H2. 379 | generalize (fun h u v => H2 h tt u tt v). generalize (H1 tt). 380 | generalize (f tt) as l. clear. 381 | intros l Hl. induction Hl as [|[h u] l H0 H1 IH]. 382 | + intros _. apply fnl_N. 383 | + intros H2. apply fnl_C. 384 | * intros H3. apply In_In_dom_lem in H3. destruct H3 as [v H4]. 385 | assert (L1: tt = tt /\ u = v). 386 | { apply (H2 h); simpl; tauto. } 387 | destruct L1 as [_ L1b]. 388 | subst v. contradiction. 389 | * { apply IH. intros k v w H3 H4. apply (H2 k). 390 | - now right. 391 | - now right. 392 | } 393 | - intros H1 H2. simpl. apply fnl_app. 394 | + apply IH. 395 | * intros bs. apply (H1 (false,bs)). 396 | * intros h bs u bs' u' H3 H4. 397 | destruct (H2 h (false,bs) u (false,bs') u' H3 H4) as [H5 H6]. 398 | split; congruence. 399 | + apply IH. 400 | * intros bs. apply (H1 (true,bs)). 401 | * intros h bs u bs' u' H3 H4. 402 | destruct (H2 h (true,bs) u (true,bs') u' H3 H4) as [H5 H6]. 403 | split; congruence. 404 | + intros h H3 H4. 405 | apply In_In_dom_lem in H3. apply In_In_dom_lem in H4. 406 | destruct H3 as [u H3]. destruct H4 as [v H4]. 407 | apply totalassets__iff in H3. apply totalassets__iff in H4. 408 | destruct H3 as [bs H3]. destruct H4 as [bs' H4]. 409 | destruct (H2 h (false,bs) u (true,bs') v H3 H4) as [H5 H6]. 410 | discriminate H5. 411 | Qed. 412 | 413 | Lemma totalassets_fnl (f:statefun) : 414 | sf_valid f -> 415 | fnl (statefun_totalassets f). 416 | intros [Hf1 [Hf2 _]]. apply totalassets_fnl_. 417 | - exact Hf1. 418 | - exact Hf2. 419 | Qed. 420 | 421 | Opaque statefun_totalassets. 422 | 423 | Lemma totalassets_trans_iff (f:statefun) (tx:Tx) fee rew : 424 | sf_valid f -> 425 | tx_valid tx -> 426 | statefun_supports_tx f tx fee rew -> 427 | forall h obl u, 428 | In (h,(obl,u)) (statefun_totalassets (tx_statefun_trans tx f)) <-> 429 | ((In (h,(obl,u)) (statefun_totalassets f) /\ ~exists alpha, In (alpha,h) (tx_inputs tx)) 430 | \/ 431 | exists alpha i, nth_error (tx_outputs tx) i = value (alpha,(obl,u)) /\ h = hashpair (hashtx tx) (hashnat i)). 432 | intros Hf Ht Hs. destruct tx as [inpl outpl]. intros h obl u. split. 433 | - intros H1. destruct (totalassets_iff (tx_statefun_trans (inpl,outpl) f) h (obl,u)) as [H2 _]. 434 | destruct (H2 H1) as [alpha H3]. 435 | unfold tx_statefun_trans in H3. apply in_app_iff in H3. 436 | destruct H3 as [H4|H4]. 437 | + apply new_assets_iff in H4. 438 | right. exists alpha. exact H4. 439 | + apply remove_assets_iff in H4. destruct H4 as [H5 H6]. left. split. 440 | * destruct (totalassets_iff f h (obl,u)) as [_ H7]. apply H7. 441 | exists alpha. assumption. 442 | * intros [beta H7]. simpl in H7. apply H6. 443 | apply get_spent_iff. simpl. 444 | assert (L1: alpha = beta). 445 | { destruct (statefun_supports_tx_assets_In f (inpl,outpl) fee rew beta h Hs H7) as [obl' [v H8]]. 446 | destruct Hf as [_ [Hf2 _]]. 447 | destruct (Hf2 h alpha (obl,u) beta (obl',v) H5 H8) as [H9 _]. 448 | exact H9. 449 | } 450 | rewrite L1. exact H7. 451 | - intros [[H1 H2]|[alpha [i H1]]]. 452 | + destruct (totalassets_iff (tx_statefun_trans (inpl,outpl) f) h (obl,u)) as [_ H3]. 453 | apply H3. 454 | destruct (totalassets_iff f h (obl,u)) as [H4 _]. 455 | destruct (H4 H1) as [alpha H5]. 456 | exists alpha. 457 | unfold tx_statefun_trans. apply in_or_app. 458 | right. apply remove_assets_iff. 459 | split. 460 | * assumption. 461 | * intros H6. apply H2. exists alpha. now apply get_spent_iff. 462 | + destruct (totalassets_iff (tx_statefun_trans (inpl,outpl) f) h (obl,u)) as [_ H3]. 463 | apply H3. 464 | exists alpha. 465 | unfold tx_statefun_trans. apply in_or_app. 466 | left. apply new_assets_iff. exists i. exact H1. 467 | Qed. 468 | 469 | Lemma sf_totalassets_app__iff {n} (f g:bitseq n -> list asset) : 470 | app_perm (totalassets_ (fun alpha => f alpha ++ g alpha)) 471 | (totalassets_ f ++ totalassets_ g). 472 | induction n as [|n IH]. 473 | - simpl. apply app_perm_ref. 474 | - simpl. 475 | generalize (IH (fun bs => f (false,bs)) (fun bs => g (false,bs))). 476 | intros IH1. 477 | generalize (IH (fun bs => f (true,bs)) (fun bs => g (true,bs))). 478 | intros IH2. 479 | apply app_perm_tra with (r := ((totalassets_ (fun bs : bitseq n => f (false, bs)) ++ 480 | totalassets_ (fun bs : bitseq n => g (false, bs))) ++ 481 | (totalassets_ (fun bs : bitseq n => f (true, bs)) ++ 482 | totalassets_ (fun bs : bitseq n => g (true, bs))))). 483 | + apply app_perm_app. 484 | * exact IH1. 485 | * exact IH2. 486 | + rewrite <- app_assoc. rewrite <- app_assoc. apply app_perm_app. 487 | * apply app_perm_ref. 488 | * { rewrite app_assoc. rewrite app_assoc. apply app_perm_app. 489 | - apply app_perm_swap. 490 | - apply app_perm_ref. 491 | } 492 | Qed. 493 | 494 | Lemma sf_totalassets__ifcons_app_perm (nw:asset) (beta:addr) {n} (alphapre:bitseq n -> addr) (f:bitseq n -> list asset) : 495 | (forall alpha1 alpha2:bitseq n, alphapre alpha1 = alphapre alpha2 -> alpha1 = alpha2) -> 496 | ((exists alpha:bitseq n, alphapre alpha = beta) -> 497 | app_perm 498 | (totalassets_ 499 | (fun alpha : bitseq n => 500 | if addr_eq_dec (alphapre alpha) beta then nw :: f alpha else f alpha)) 501 | (nw :: totalassets_ f)) 502 | /\ 503 | (~(exists alpha:bitseq n, alphapre alpha = beta) -> 504 | app_perm 505 | (totalassets_ 506 | (fun alpha : bitseq n => 507 | if addr_eq_dec (alphapre alpha) beta then nw :: f alpha else f alpha)) 508 | (totalassets_ f)). 509 | induction n as [|n IH]; split; intros H1. 510 | - simpl. destruct H1 as [[] H1]. 511 | destruct (addr_eq_dec (alphapre tt) beta) as [H2|H2]. 512 | + apply app_perm_ref. 513 | + exfalso. contradiction. 514 | - simpl. 515 | destruct (addr_eq_dec (alphapre tt) beta) as [H2|H2]. 516 | + exfalso. apply H1. exists tt. assumption. 517 | + apply app_perm_ref. 518 | - set (alphapre1 := fun alpha:bitseq n => alphapre (true,alpha)). 519 | set (alphapre0 := fun alpha:bitseq n => alphapre (false,alpha)). 520 | assert (alphapre1inj : forall alpha1 alpha2, alphapre1 alpha1 = alphapre1 alpha2 -> alpha1 = alpha2). 521 | { intros alpha1 alpha2 H2. apply H in H2. congruence. } 522 | assert (alphapre0inj : forall alpha1 alpha2, alphapre0 alpha1 = alphapre0 alpha2 -> alpha1 = alpha2). 523 | { intros alpha1 alpha2 H2. apply H in H2. congruence. } 524 | destruct H1 as [[[|] alpha] H1]. 525 | + destruct (IH alphapre1 (fun alpha => f (true,alpha)) alphapre1inj) as [IH1 _]. 526 | destruct (IH alphapre0 (fun alpha => f (false,alpha)) alphapre0inj) as [_ IH2]. 527 | simpl. 528 | apply app_perm_tra with (r := (totalassets_ (fun bs : bitseq n => f (false, bs)) ++ nw::totalassets_ (fun bs : bitseq n => f (true, bs)))). 529 | * { apply app_perm_app. 530 | - apply IH2. intros [alpha0 H2]. rewrite <- H1 in H2. apply H in H2. 531 | discriminate H2. 532 | - apply IH1. exists alpha. exact H1. 533 | } 534 | * { change (app_perm 535 | (totalassets_ (fun bs : bitseq n => f (false, bs)) ++ 536 | ((cons nw nil) ++ totalassets_ (fun bs : bitseq n => f (true, bs)))) 537 | ((nw 538 | :: totalassets_ (fun bs : bitseq n => f (false, bs))) ++ 539 | totalassets_ (fun bs : bitseq n => f (true, bs)))). 540 | rewrite app_assoc. apply app_perm_app. 541 | - apply (app_perm_swap _ (cons nw nil)). 542 | - apply app_perm_ref. 543 | } 544 | + destruct (IH alphapre1 (fun alpha => f (true,alpha)) alphapre1inj) as [_ IH1]. 545 | destruct (IH alphapre0 (fun alpha => f (false,alpha)) alphapre0inj) as [IH2 _]. 546 | simpl. 547 | change (app_perm 548 | (totalassets_ 549 | (fun bs : bitseq n => 550 | if addr_eq_dec (alphapre (false, bs)) beta 551 | then nw :: f (false, bs) 552 | else f (false, bs)) ++ 553 | totalassets_ 554 | (fun bs : bitseq n => 555 | if addr_eq_dec (alphapre (true, bs)) beta 556 | then nw :: f (true, bs) 557 | else f (true, bs))) 558 | ((nw 559 | :: totalassets_ (fun bs : bitseq n => f (false, bs))) ++ 560 | totalassets_ (fun bs : bitseq n => f (true, bs)))). 561 | apply app_perm_app. 562 | * apply IH2. exists alpha. exact H1. 563 | * apply IH1. intros [alpha0 H2]. rewrite <- H1 in H2. apply H in H2. 564 | discriminate H2. 565 | - set (alphapre1 := fun alpha:bitseq n => alphapre (true,alpha)). 566 | set (alphapre0 := fun alpha:bitseq n => alphapre (false,alpha)). 567 | assert (alphapre1inj : forall alpha1 alpha2, alphapre1 alpha1 = alphapre1 alpha2 -> alpha1 = alpha2). 568 | { intros alpha1 alpha2 H2. apply H in H2. congruence. } 569 | assert (alphapre0inj : forall alpha1 alpha2, alphapre0 alpha1 = alphapre0 alpha2 -> alpha1 = alpha2). 570 | { intros alpha1 alpha2 H2. apply H in H2. congruence. } 571 | destruct (IH alphapre1 (fun alpha => f (true,alpha)) alphapre1inj) as [_ IH1]. 572 | destruct (IH alphapre0 (fun alpha => f (false,alpha)) alphapre0inj) as [_ IH2]. 573 | simpl. apply app_perm_app. 574 | + apply IH2. intros [alpha H2]. apply H1. exists (false,alpha). exact H2. 575 | + apply IH1. intros [alpha H2]. apply H1. exists (true,alpha). exact H2. 576 | Qed. 577 | 578 | Lemma sf_totalassets__remove_assets {n} (alphapre:bitseq n -> addr) (f:bitseq n -> list asset) (inpl:list addr_assetid) : 579 | (forall h alpha u beta v, In (h,u) (f alpha) -> In (h,v) (f beta) -> alpha = beta /\ u = v) -> 580 | (forall h alpha, In (alphapre alpha,h) inpl -> exists u, In (h,u) (f alpha)) -> 581 | (forall h u, In (h,u) (totalassets_ f) -> (In h (map (@snd addr hashval) inpl) <-> exists alpha, In (alphapre alpha,h) inpl)) -> 582 | app_perm (remove_assets (totalassets_ f) (map (@snd addr hashval) inpl)) 583 | (totalassets_ (fun alpha => remove_assets (f alpha) (get_spent (alphapre alpha) inpl))). 584 | intros Hf2 Ht1. induction n as [|n IH]. 585 | - simpl. 586 | intros H1. rewrite (remove_assets_eq (f tt) (map (snd (B:=hashval)) inpl) (get_spent (alphapre tt) inpl)). 587 | + apply app_perm_ref. 588 | + intros h H2. apply In_In_dom_lem in H2. destruct H2 as [u H2]. split. 589 | * intros H3. apply (H1 _ _ H2) in H3. destruct H3 as [[] H3]. now apply get_spent_iff. 590 | * intros H3. apply (H1 _ _ H2). exists tt. now apply get_spent_iff. 591 | - simpl. intros H1. 592 | apply app_perm_tra with (r := 593 | (remove_assets 594 | (totalassets_ (fun bs : bitseq n => f (false, bs))) 595 | (map (snd (B:=hashval)) inpl)) 596 | ++ 597 | (remove_assets 598 | (totalassets_ (fun bs : bitseq n => f (true, bs))) 599 | (map (snd (B:=hashval)) inpl))). 600 | + rewrite remove_assets_app. apply app_perm_ref. 601 | + apply app_perm_app. 602 | * { set (alphapre0 := fun alpha:bitseq n => alphapre (false,alpha)). 603 | set (f0 := fun alpha:bitseq n => f (false,alpha)). 604 | apply (IH alphapre0 f0). 605 | - intros h alpha u beta v H2 H3. 606 | destruct (Hf2 h (false,alpha) u (false,beta) v H2 H3) as [H4 H5]. 607 | split; congruence. 608 | - intros h alpha H2. unfold alphapre0 in H2. apply Ht1 in H2. exact H2. 609 | - intros h u H2. split. 610 | + intros H3. apply (H1 h u) in H3. 611 | * { destruct H3 as [[[|] alpha] H4]. 612 | - apply totalassets__iff in H2. unfold f0 in H2. simpl in H2. 613 | destruct H2 as [beta H2]. 614 | apply Ht1 in H4. destruct H4 as [v H5]. 615 | destruct (Hf2 h (true,alpha) v (false,beta) u H5 H2) as [H6 _]. 616 | discriminate H6. 617 | - exists alpha. exact H4. 618 | } 619 | * apply in_or_app. left. exact H2. 620 | + intros [alpha H3]. apply (H1 h u). 621 | * apply in_or_app. left. exact H2. 622 | * exists (false,alpha). exact H3. 623 | } 624 | * { set (alphapre1 := fun alpha:bitseq n => alphapre (true,alpha)). 625 | set (f1 := fun alpha:bitseq n => f (true,alpha)). 626 | apply (IH alphapre1 f1). 627 | - intros h alpha u beta v H2 H3. 628 | destruct (Hf2 h (true,alpha) u (true,beta) v H2 H3) as [H4 H5]. 629 | split; congruence. 630 | - intros h alpha H2. unfold alphapre1 in H2. apply Ht1 in H2. exact H2. 631 | - intros h u H2. split. 632 | + intros H3. apply (H1 h u) in H3. 633 | * { destruct H3 as [[[|] alpha] H4]. 634 | - exists alpha. exact H4. 635 | - apply totalassets__iff in H2. unfold f1 in H2. simpl in H2. 636 | destruct H2 as [beta H2]. 637 | apply Ht1 in H4. destruct H4 as [v H5]. 638 | destruct (Hf2 h (false,alpha) v (true,beta) u H5 H2) as [H6 _]. 639 | discriminate H6. 640 | } 641 | * apply in_or_app. right. exact H2. 642 | + intros [alpha H3]. apply (H1 h u). 643 | * apply in_or_app. right. exact H2. 644 | * exists (true,alpha). exact H3. 645 | } 646 | Qed. 647 | 648 | Transparent statefun_totalassets. 649 | 650 | Lemma sf_totalassets_app_iff f g : 651 | app_perm (statefun_totalassets (fun alpha => f alpha ++ g alpha)) 652 | (statefun_totalassets f ++ statefun_totalassets g). 653 | apply sf_totalassets_app__iff. 654 | Qed. 655 | 656 | Lemma sf_totalassets_ifcons_app_perm beta nw f : 657 | app_perm (statefun_totalassets (fun alpha => if addr_eq_dec alpha beta then nw::f alpha else f alpha)) (nw::statefun_totalassets f). 658 | apply sf_totalassets__ifcons_app_perm. 659 | - tauto. 660 | - exists beta. reflexivity. 661 | Qed. 662 | 663 | Lemma sf_totalassets_remove_assets f inpl : 664 | (forall h alpha u beta v, In (h,u) (f alpha) -> In (h,v) (f beta) -> alpha = beta /\ u = v) -> 665 | (forall h alpha, In (alpha,h) inpl -> exists u, In (h,u) (f alpha)) -> 666 | app_perm (remove_assets (statefun_totalassets f) (map (@snd addr hashval) inpl)) 667 | (statefun_totalassets (fun alpha => remove_assets (f alpha) (get_spent alpha inpl))). 668 | intros H1 H2. 669 | apply (sf_totalassets__remove_assets (fun alpha => alpha) f inpl H1 H2). 670 | intros h u H3. split. 671 | - intros H4. apply in_map_iff in H4. destruct H4 as [[alpha h'] [H5 H6]]. 672 | exists alpha. simpl in H5. congruence. 673 | - intros [alpha H4]. apply in_map_iff. exists (alpha,h). split. 674 | + reflexivity. 675 | + exact H4. 676 | Qed. 677 | 678 | Opaque statefun_totalassets. 679 | 680 | Lemma sf_totalassets_nil : statefun_totalassets (fun alpha => nil) = nil. 681 | assert (L1: exists l, statefun_totalassets (fun alpha => nil) = l). 682 | { now exists (statefun_totalassets (fun alpha => nil)). } 683 | destruct L1 as [[|[h u] l] L1a]. 684 | - exact L1a. 685 | - assert (L2: In (h, u) (statefun_totalassets (fun _ => nil))). 686 | { rewrite L1a. left. reflexivity. } 687 | destruct (totalassets_iff (fun _ => nil) h u) as [H1 H2]. 688 | destruct (H1 L2) as [alpha []]. 689 | Qed. 690 | 691 | Opaque asset_value_sum. 692 | 693 | Lemma sf_totalunits_new_assets (aul:list addr_preasset) txh i : 694 | statefun_totalunits (fun alpha => new_assets alpha aul txh i) = 695 | asset_value_out aul. 696 | revert i. induction aul as [|[beta [obl u]] aul IH]; intros i. 697 | - simpl. unfold statefun_totalunits. rewrite sf_totalassets_nil. simpl. 698 | reflexivity. 699 | - simpl. rewrite <- (IH (S i)). unfold statefun_totalunits. 700 | assert (L1: app_perm (statefun_totalassets 701 | (fun alpha : addr => 702 | if addr_eq_dec alpha beta 703 | then 704 | (hashpair txh (hashnat i), (obl,u)) :: new_assets alpha aul txh (S i) 705 | else new_assets alpha aul txh (S i))) 706 | ((hashpair txh (hashnat i),(obl,u))::(statefun_totalassets 707 | (fun alpha : addr => new_assets alpha aul txh (S i))))). 708 | { apply (sf_totalassets_ifcons_app_perm beta (hashpair txh (hashnat i),(obl,u))). } 709 | apply app_perm_asset_value_sum in L1. 710 | simpl in L1. rewrite L1. 711 | destruct (asset_value (hashpair txh (hashnat i), (obl,u))) as [u'|] eqn:E1. 712 | + rewrite (asset_value_sum_value_cons _ _ _ E1). 713 | destruct u as [u|d]; try discriminate E1. 714 | unfold asset_value in E1. simpl in E1. inversion E1. reflexivity. 715 | + rewrite (asset_value_sum_novalue_cons _ _ E1). 716 | destruct u as [n|d]; try reflexivity; discriminate E1. 717 | Qed. 718 | 719 | Lemma no_dups_map_inpl f inpl : 720 | sf_valid f -> 721 | (forall h alpha, In (alpha,h) inpl -> exists u, In (h,u) (f alpha)) -> 722 | no_dups inpl -> 723 | no_dups (map (@snd addr hashval) inpl). 724 | intros [_ [Hf2 _]] H1 H. 725 | induction H as [|[alpha h] inpl H2 H3 IH]. 726 | - apply no_dups_nil. 727 | - simpl. apply no_dups_cons. 728 | + intros H4. apply in_map_iff in H4. 729 | destruct H4 as [[beta k] [H5 H6]]. simpl in H5. subst k. 730 | assert (L1: exists oblu : prod obligation preasset, In (h, oblu) (f alpha)). 731 | { apply H1. now left. } 732 | assert (L2: exists oblu : prod obligation preasset, In (h, oblu) (f beta)). 733 | { apply H1. now right. } 734 | destruct L1 as [[obl u] L1a]. 735 | destruct L2 as [[obl' v] L2a]. 736 | destruct (Hf2 h alpha (obl,u) beta (obl',v) L1a L2a) as [H7 _]. 737 | subst beta. contradiction. 738 | + apply IH. intros k beta H4. apply H1. now right. 739 | Qed. 740 | 741 | Lemma statefun_asset_value_in_lem f inpl utot : 742 | sf_valid f -> 743 | (forall h alpha, In (alpha,h) inpl -> exists u, In (h,u) (f alpha)) -> 744 | no_dups inpl -> 745 | statefun_asset_value_in f inpl utot -> 746 | utot + statefun_totalunits 747 | (fun alpha : addr => 748 | remove_assets (f alpha) (get_spent alpha inpl)) 749 | = 750 | statefun_totalunits f. 751 | intros Hf Ht1 Ht3 H. 752 | assert (L1:forall (h : hashval) (alpha : addr) (u : prod obligation preasset) 753 | (beta : addr) (v : prod obligation preasset), 754 | In (h, u) (f alpha) -> In (h, v) (f beta) -> alpha = beta /\ u = v). 755 | { destruct Hf as [_ [Hf2 _]]. exact Hf2. } 756 | unfold statefun_totalunits at 1. 757 | rewrite <- (app_perm_asset_value_sum _ _ (sf_totalassets_remove_assets f inpl L1 Ht1)). 758 | clear L1. 759 | induction H as [|h a u inpl alpha v H1 IH1 H2|h a inpl alpha v H1 IH1 H2 H3]. 760 | - simpl. rewrite remove_assets_nil. unfold statefun_totalunits. omega. 761 | - simpl. 762 | assert (L3:forall h alpha, In (alpha, h) inpl -> exists u : prod obligation preasset, In (h, u) (f alpha)). 763 | { intros k beta H3. apply Ht1. now right. } 764 | specialize (IH1 L3). 765 | assert (L4: v + (u + asset_value_sum 766 | (remove_assets (statefun_totalassets f) 767 | (h :: map (snd (B:=hashval)) inpl))) 768 | = statefun_totalunits f). 769 | { rewrite asset_value_sum_asset_shift with (a := (assetobl a,assetpre a)). 770 | - apply IH1. inversion Ht3. assumption. 771 | - apply totalassets_fnl. exact Hf. 772 | - exact (no_dups_map_inpl f _ Hf Ht1 Ht3). 773 | - apply totalassets_iff. exists alpha. subst h. 774 | destruct a as [h [obl u']]. exact H2. 775 | - exact H. 776 | } 777 | omega. 778 | - simpl. 779 | assert (L3:forall h alpha, In (alpha, h) inpl -> exists u : prod obligation preasset, In (h, u) (f alpha)). 780 | { intros k beta H4. apply Ht1. now right. } 781 | specialize (IH1 L3). 782 | assert (L4: v + asset_value_sum 783 | (remove_assets (statefun_totalassets f) 784 | (h :: map (snd (B:=hashval)) inpl)) 785 | = statefun_totalunits f). 786 | { rewrite asset_value_sum_asset_shift_non with (a := (assetobl a,assetpre a)). 787 | - apply IH1. inversion Ht3. assumption. 788 | - apply totalassets_fnl. exact Hf. 789 | - exact (no_dups_map_inpl f _ Hf Ht1 Ht3). 790 | - apply totalassets_iff. exists alpha. subst h. 791 | destruct a as [h [obl u']]. exact H2. 792 | - exact H3. 793 | } 794 | omega. 795 | Qed. 796 | 797 | Theorem totalunits_bdd (f:statefun) (tx:Tx) (fee rew:nat) : 798 | sf_valid f -> 799 | tx_valid tx -> 800 | statefun_supports_tx f tx fee rew -> 801 | statefun_totalunits (tx_statefun_trans tx f) + fee = statefun_totalunits f + rew. 802 | intros Hf Ht Hs. generalize Hs. intros [utot [H1 H2]]. 803 | unfold statefun_totalunits at 1. unfold tx_statefun_trans. 804 | generalize (sf_totalassets_app_iff (fun alpha => new_assets alpha (tx_outputs tx) (hashtx tx) 0) (fun alpha => remove_assets (f alpha) (get_spent alpha (tx_inputs tx)))). 805 | intros H3. apply app_perm_asset_value_sum in H3. 806 | rewrite H3. clear H3. 807 | rewrite <- (asset_value_sum_app 808 | (statefun_totalassets 809 | (fun alpha : addr => 810 | new_assets alpha (tx_outputs tx) (hashtx tx) 0)) 811 | (statefun_totalassets 812 | (fun alpha : addr => 813 | remove_assets (f alpha) (get_spent alpha (tx_inputs tx))))). 814 | change ((statefun_totalunits 815 | (fun alpha : addr => 816 | new_assets alpha (tx_outputs tx) (hashtx tx) 0)) + 817 | (statefun_totalunits 818 | (fun alpha : addr => 819 | remove_assets (f alpha) (get_spent alpha (tx_inputs tx)))) + fee = 820 | statefun_totalunits f + rew). 821 | rewrite sf_totalunits_new_assets. 822 | assert (L1:utot + (statefun_totalunits 823 | (fun alpha : addr => 824 | remove_assets (f alpha) (get_spent alpha (tx_inputs tx)))) = 825 | statefun_totalunits f). 826 | { apply statefun_asset_value_in_lem. 827 | - exact Hf. 828 | - destruct tx as [inpl outpl]. destruct Ht as [Ht1a _]. 829 | simpl. intros h alpha H3. 830 | destruct (statefun_supports_tx_assets_In f (inpl,outpl) fee rew alpha h Hs H3) as [obl' [v H4]]. 831 | exists (obl',v). exact H4. 832 | - destruct tx as [inpl outpl]. destruct Ht as [_ Ht2b]. exact Ht2b. 833 | - exact H1. 834 | } 835 | omega. 836 | Qed. 837 | 838 | -------------------------------------------------------------------------------- /Prelude.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** Prelude: Import parts of Coq's library and add a few extras. **) 6 | 7 | Require Export Coq.Lists.List. 8 | Require Export Coq.Arith.Peano_dec. 9 | Require Export omega.Omega. 10 | 11 | Lemma nth_error_In {X:Type} (l:list X) (i:nat) (x:X) : 12 | nth_error l i = value x -> In x l. 13 | revert i. induction l as [|y l IH]; intros [|i]. 14 | - intros H1. discriminate H1. 15 | - intros H1. discriminate H1. 16 | - intros H1. inversion H1. now left. 17 | - intros H1. right. apply (IH i). exact H1. 18 | Qed. 19 | 20 | Lemma nth_error_app {X:Type} (l r:list X) (i:nat) (x:X) : 21 | nth_error (l ++ r) i = value x -> 22 | nth_error l i = value x \/ exists j, i = length l + j /\ nth_error r j = value x. 23 | revert i. induction l as [|y l IH]. 24 | - simpl. intros i H1. right. exists i. tauto. 25 | - intros [|i]; simpl. 26 | + intros H1. now left. 27 | + intros H1. destruct (IH i H1) as [H2|[j [H2 H3]]]. 28 | * now left. 29 | * { right. exists j. split. 30 | - congruence. 31 | - exact H3. 32 | } 33 | Qed. 34 | 35 | Fixpoint In_dom {X Y:Type} (x : X) (xyl : list (X * Y)) : Prop := 36 | match xyl with 37 | | nil => False 38 | | cons (x',_) xyl' => x = x' \/ In_dom x xyl' 39 | end. 40 | 41 | Lemma In_In_dom_lem {X Y:Type} (x : X) (xyl : list (X * Y)) : 42 | In_dom x xyl <-> exists y:Y, In (x,y) xyl. 43 | induction xyl as [|[x' y'] l' [IH1 IH2]]. 44 | - simpl. split. 45 | + tauto. 46 | + intros [y []]. 47 | - simpl. split. 48 | + intros [H1|H1]. 49 | * subst x. exists y'. tauto. 50 | * destruct (IH1 H1) as [y H2]. exists y. tauto. 51 | + intros [y [H1|H1]]. 52 | * inversion H1. tauto. 53 | * right. apply IH2. exists y. assumption. 54 | Qed. 55 | 56 | Lemma In_In_dom_lem_2 {X Y:Type} (x : X) (y : Y) (xyl : list (X * Y)) : 57 | In (x,y) xyl -> In_dom x xyl. 58 | intros H. apply In_In_dom_lem. exists y. exact H. 59 | Qed. 60 | 61 | Inductive no_dups {X:Type} : list X -> Prop := 62 | | no_dups_nil : no_dups nil 63 | | no_dups_cons x l : ~ In x l -> no_dups l -> no_dups (x::l) 64 | . 65 | 66 | Inductive fnl {X Y:Type} : list (X * Y) -> Prop := 67 | | fnl_N : fnl nil 68 | | fnl_C x y l : ~ In_dom x l -> fnl l -> fnl (cons (x,y) l) 69 | . 70 | 71 | Lemma fnl_lem {X Y:Type} (l:list (X * Y)) : 72 | fnl l -> forall x y z, In (x,y) l -> In (x,z) l -> y = z. 73 | intros H. induction H as [|w v l H1 H2 IH]. 74 | - intros x y z []. 75 | - intros x y z [H3|H3] [H4|H4]. 76 | + inversion H3. inversion H4. congruence. 77 | + exfalso. apply In_In_dom_lem_2 in H4. inversion H3. congruence. 78 | + exfalso. apply In_In_dom_lem_2 in H3. inversion H4. congruence. 79 | + revert H3 H4. apply IH. 80 | Qed. 81 | 82 | Inductive app_perm {X:Type} : list X -> list X -> Prop := 83 | | app_perm_swap l r : app_perm (l ++ r) (r ++ l) 84 | | app_perm_app l1 r1 l2 r2 : app_perm l1 r1 -> app_perm l2 r2 -> app_perm (l1 ++ l2) (r1 ++ r2) 85 | | app_perm_ref l : app_perm l l 86 | | app_perm_sym l r : app_perm l r -> app_perm r l 87 | | app_perm_tra l r s : app_perm l r -> app_perm r s -> app_perm l s 88 | . 89 | 90 | Lemma no_dups_app {X:Type} (l r:list X) : 91 | no_dups l -> no_dups r -> (forall x, In x l -> ~In x r) -> no_dups (l ++ r). 92 | intros H H1. induction H as [|x l H2 H3 IH]. 93 | - simpl. tauto. 94 | - intros H4. simpl. apply no_dups_cons. 95 | + intros H5. apply in_app_iff in H5. destruct H5 as [H5|H5]. 96 | * contradiction. 97 | * revert H5. apply H4. now left. 98 | + apply IH. intros y H5. apply H4. now right. 99 | Qed. 100 | 101 | Lemma fnl_app {X Y:Type} (l r:list (prod X Y)) : 102 | fnl l -> fnl r -> (forall x, In_dom x l -> ~In_dom x r) -> fnl (l ++ r). 103 | intros H H1. induction H as [|x y l H2 H3 IH]. 104 | - simpl. tauto. 105 | - intros H4. simpl. apply fnl_C. 106 | + intros H5. apply In_In_dom_lem in H5. destruct H5 as [z H5]. 107 | apply in_app_iff in H5. destruct H5 as [H5|H5]. 108 | * apply H2. now apply (In_In_dom_lem_2 x z). 109 | * apply In_In_dom_lem_2 in H5. revert H5. apply H4. now left. 110 | + apply IH. intros z H5. apply H4. now right. 111 | Qed. 112 | 113 | 114 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # ledgertheory 2 | 3 | ## Ledger Theory Coq Development 4 | 5 | This repository contains a Coq development of a theory of lightweight cryptographic ledgers. 6 | The files should be compiled in the following order (see also the file coqcompile): 7 | 8 | Prelude.v 9 | Addrs.v 10 | CryptoHashes.v 11 | CryptoSignatures.v 12 | Assets.v 13 | Transactions.v 14 | LedgerStates.v 15 | MTrees.v 16 | CTrees.v 17 | CTreeGrafting.v 18 | Blocks.v 19 | 20 | ## White Paper 21 | 22 | There is also a white paper lightcrypto.pdf providing a high level description and overview 23 | of what is here. More detailed descriptions of the various parts will hopefully be written soon. 24 | 25 | ## License 26 | 27 | This is all being released as open source under the MIT/X11 software license. 28 | See http://www.opensource.org/licenses/mit-license.php 29 | This is also in the file LICENSE along with a bitcoin signature of the license. 30 | -------------------------------------------------------------------------------- /Transactions.v: -------------------------------------------------------------------------------- 1 | (** Copyright (c) 2015 Bill White **) 2 | (** Distributed under the MIT/X11 software license **) 3 | (** See http://www.opensource.org/licenses/mit-license.php **) 4 | 5 | (** Transactions: Transactions are represented as pairs of input and output lists. 6 | An input list contains pairs of addresses and hashvals referring to an asset by 7 | it's id and the address where the asset is held. 8 | An output list contains addresses associated with pairs of use requirements and preassets. 9 | An asset id h is constructed for each output by hashing the transaction along with the index 10 | of the output (vout). Then the corresponding asset will be constructed by 11 | combining the asset id with the use requirement and preasset. 12 | A transaction is considered valid if it has at least one input and no duplicate inputs. 13 | This should not be confused with the transaction being supported, a notion 14 | that depends on the ledger. 15 | **) 16 | 17 | Require Export Assets. 18 | 19 | (*** 20 | Basic unsigned Tx consisting of inputs and outputs 21 | where the input assets are only referred by address and hashval (assetid) 22 | ***) 23 | Definition Tx : Type := prod (list addr_assetid) (list addr_preasset). 24 | 25 | Definition hashtx (tx : Tx) : hashval := 26 | match tx with 27 | | (inpl,outpl) => 28 | hashpair (hashlist (map hash_addr_assetid inpl)) 29 | (hashlist (map hash_addr_preasset outpl)) 30 | end. 31 | 32 | Definition tx_inputs (tx:Tx) : list addr_assetid := 33 | match tx with 34 | | (inpl,outpl) => inpl 35 | end. 36 | 37 | Definition tx_outputs (tx:Tx) : list addr_preasset := 38 | match tx with 39 | | (inpl,outpl) => outpl 40 | end. 41 | 42 | Lemma hashtxioinj inpl outpl inpl' outpl' : 43 | hashtx (inpl,outpl) = hashtx (inpl',outpl') 44 | -> inpl = inpl' /\ outpl = outpl'. 45 | unfold hashtx. intros H1. 46 | apply hashpairinj in H1. destruct H1 as [H2 H3]. split. 47 | - apply hashlistinj in H2. now apply map_hash_addr_assetidinj. 48 | - apply hashlistinj in H3. now apply map_hash_addr_preassetinj. 49 | Qed. 50 | 51 | Lemma hashtx_notin_inpl_lem alpha inpl inpl' outpl i : 52 | (forall z, In z inpl' -> In z inpl) -> 53 | ~In (alpha,hashpair (hashtx (inpl,outpl)) (hashnat i)) inpl'. 54 | induction inpl'. 55 | - simpl; tauto. 56 | - intros H1 [H2|H2]. 57 | + destruct a as [gamma hc]. 58 | inversion H2. 59 | apply (subh_irrefl hc). 60 | rewrite H3 at 2. 61 | apply subh_PL. apply subh_PL. 62 | apply subh_tra with (h2 := (hash_addr_assetid (gamma,hc))). 63 | * simpl. apply subh_R. 64 | * apply subh_hashlist. apply in_map. apply H1. now left. 65 | + revert H2. apply IHinpl'. 66 | intros z H3. apply H1. now right. 67 | Qed. 68 | 69 | Lemma hashtx_notin_inpl alpha inpl outpl i : 70 | ~In (alpha,hashpair (hashtx (inpl,outpl)) (hashnat i)) inpl. 71 | apply hashtx_notin_inpl_lem. 72 | tauto. 73 | Qed. 74 | 75 | Inductive inpl_reln {n:nat} (alphapre:bitseq n -> addr) : list addr_assetid -> list (bitseq n * hashval) -> Prop := 76 | | inpl_reln_nil : inpl_reln alphapre nil nil 77 | | inpl_reln_skip fullinpl inpl alpha h : 78 | (forall gamma:bitseq n, alphapre gamma <> alpha) 79 | -> inpl_reln alphapre fullinpl inpl 80 | -> inpl_reln alphapre ((alpha,h)::fullinpl) inpl 81 | | inpl_reln_cons fullinpl inpl gamma h : 82 | inpl_reln alphapre fullinpl inpl 83 | -> inpl_reln alphapre ((alphapre gamma,h)::fullinpl) ((gamma,h)::inpl) 84 | . 85 | 86 | Inductive outpl_reln (txh:hashval) {n:nat} (alphapre:bitseq n -> addr) : nat -> list addr_preasset -> list (bitseq n * asset) -> Prop := 87 | | outpl_reln_nil i : outpl_reln txh alphapre i nil nil 88 | | outpl_reln_skip i fulloutpl outpl alpha u : 89 | (forall gamma:bitseq n, alphapre gamma <> alpha) 90 | -> outpl_reln txh alphapre (S i) fulloutpl outpl 91 | -> outpl_reln txh alphapre i ((alpha,u)::fulloutpl) outpl 92 | | outpl_reln_cons i fulloutpl outpl gamma u : 93 | outpl_reln txh alphapre (S i) fulloutpl outpl 94 | -> outpl_reln txh alphapre i ((alphapre gamma,u)::fulloutpl) ((gamma,(hashpair txh (hashnat i),u))::outpl) 95 | . 96 | 97 | Lemma inpl_reln_start fullinpl : 98 | inpl_reln (fun alpha : addr => alpha) fullinpl fullinpl. 99 | induction fullinpl as [|[alpha h] fullinpl IH]. 100 | - simpl. apply inpl_reln_nil. 101 | - simpl. 102 | apply (inpl_reln_cons (fun alpha:addr => alpha) fullinpl fullinpl). 103 | apply IH. 104 | Qed. 105 | 106 | Lemma outpl_reln_start txh fulloutpl j : 107 | outpl_reln txh (fun alpha : addr => alpha) j fulloutpl (add_vout txh fulloutpl j). 108 | revert j. induction fulloutpl as [|[alpha [obl u]] fulloutpl IH]. 109 | - simpl. apply outpl_reln_nil. 110 | - simpl. intros j. 111 | apply (outpl_reln_cons txh (fun alpha:addr => alpha) j fulloutpl (add_vout txh fulloutpl (S j)) alpha (obl,u)). 112 | apply IH. 113 | Qed. 114 | 115 | Lemma inpl_reln_In_iff {n:nat} (alphapre:bitseq n -> addr) fullinpl inpl : 116 | (forall gamma delta, alphapre gamma = alphapre delta -> gamma = delta) -> 117 | inpl_reln alphapre fullinpl inpl -> 118 | forall gamma h, In (gamma,h) inpl <-> In (alphapre gamma,h) fullinpl. 119 | intros H0 H. induction H as [|fullinpl inpl alpha k H1 H2 IH|fullinpl inpl delta k H1 IH]. 120 | - simpl. tauto. 121 | - simpl. intros gamma h. split. 122 | + intros H3. right. now apply IH. 123 | + intros [H3|H3]. 124 | * exfalso. inversion H3. symmetry in H4. revert H4. apply H1. 125 | * now apply IH. 126 | - intros gamma h. split. 127 | + intros [H3|H3]. 128 | * inversion H3. now left. 129 | * right. now apply IH. 130 | + intros [H3|H3]. 131 | * inversion H3. apply H0 in H2. left. congruence. 132 | * right. now apply IH. 133 | Qed. 134 | 135 | Lemma outpl_reln_new_assets_eq1 (fulloutpl:list addr_preasset) (txh:hashval) : 136 | forall j, forall (outpl:list (bitseq 0 * asset)%type), 137 | forall (alphapre:bitseq 0 -> addr), 138 | outpl_reln txh alphapre j fulloutpl outpl -> 139 | new_assets (alphapre tt) fulloutpl txh j = map (snd (B:=asset)) outpl. 140 | intros j outpl alphapre H. 141 | induction H as [j|j fulloutpl outpl alpha [obl u] H1 H2 IH|j fulloutpl outpl [] [obl u] H1 IH]. 142 | - simpl. reflexivity. 143 | - simpl. destruct (addr_eq_dec (alphapre tt) alpha) as [H3|H3]. 144 | + exfalso. apply (H1 tt). exact H3. 145 | + exact IH. 146 | - simpl. destruct (addr_eq_dec (alphapre tt) (alphapre tt)) as [H3|H3]. 147 | + rewrite IH. reflexivity. 148 | + exfalso. apply H3. reflexivity. 149 | Qed. 150 | 151 | Lemma inpl_reln_strip_bitseq_false {n:nat} 152 | (alphapre:bitseq (S n) -> addr) (fullinpl:list addr_assetid) (inpl:list (bitseq (S n) * hashval)%type) : 153 | (forall gamma delta, alphapre gamma = alphapre delta -> gamma = delta) -> 154 | inpl_reln alphapre fullinpl inpl -> 155 | inpl_reln (fun gamma => alphapre (false,gamma)) fullinpl (strip_bitseq_false inpl). 156 | intros Hapi H. induction H as [|fullinpl inpl alpha h H1 H2 IH|fullinpl inpl [[|] gamma] h H1 IH]. 157 | - simpl. apply inpl_reln_nil. 158 | - simpl. apply inpl_reln_skip. 159 | + intros gamma. apply H1. 160 | + exact IH. 161 | - simpl. apply inpl_reln_skip. 162 | + intros gamma' H2. apply Hapi in H2. inversion H2. 163 | + exact IH. 164 | - simpl. 165 | apply (inpl_reln_cons (fun gamma => alphapre (false,gamma)) fullinpl (strip_bitseq_false inpl) gamma h). 166 | exact IH. 167 | Qed. 168 | 169 | Lemma inpl_reln_strip_bitseq_true {n:nat} 170 | (alphapre:bitseq (S n) -> addr) (fullinpl:list addr_assetid) (inpl:list (bitseq (S n) * hashval)%type) : 171 | (forall gamma delta, alphapre gamma = alphapre delta -> gamma = delta) -> 172 | inpl_reln alphapre fullinpl inpl -> 173 | inpl_reln (fun gamma => alphapre (true,gamma)) fullinpl (strip_bitseq_true inpl). 174 | intros Hapi H. induction H as [|fullinpl inpl alpha h H1 H2 IH|fullinpl inpl [[|] gamma] h H1 IH]. 175 | - simpl. apply inpl_reln_nil. 176 | - simpl. apply inpl_reln_skip. 177 | + intros gamma. apply H1. 178 | + exact IH. 179 | - simpl. 180 | apply (inpl_reln_cons (fun gamma => alphapre (true,gamma)) fullinpl (strip_bitseq_true inpl) gamma h). 181 | exact IH. 182 | - simpl. apply inpl_reln_skip. 183 | + intros gamma' H2. apply Hapi in H2. inversion H2. 184 | + exact IH. 185 | Qed. 186 | 187 | Lemma outpl_reln_strip_bitseq_false (txh:hashval) {n:nat} 188 | (alphapre:bitseq (S n) -> addr) (fulloutpl:list addr_preasset) (outpl:list (bitseq (S n) * asset)%type) : 189 | (forall gamma delta, alphapre gamma = alphapre delta -> gamma = delta) -> 190 | forall j, 191 | outpl_reln txh alphapre j fulloutpl outpl -> 192 | outpl_reln txh (fun gamma => alphapre (false,gamma)) j fulloutpl (strip_bitseq_false outpl). 193 | intros Hapi j H. induction H as [j|j fulloutpl outpl alpha u H1 H2 IH|j fulloutpl outpl [[|] gamma] u H1 IH]. 194 | - simpl. apply outpl_reln_nil. 195 | - simpl. apply outpl_reln_skip. 196 | + intros gamma. apply H1. 197 | + exact IH. 198 | - simpl. apply outpl_reln_skip. 199 | + intros gamma' H2. apply Hapi in H2. inversion H2. 200 | + exact IH. 201 | - simpl. 202 | apply (outpl_reln_cons txh (fun gamma => alphapre (false,gamma)) j fulloutpl (strip_bitseq_false outpl) gamma u). 203 | exact IH. 204 | Qed. 205 | 206 | Lemma outpl_reln_strip_bitseq_true (txh:hashval) {n:nat} 207 | (alphapre:bitseq (S n) -> addr) (fulloutpl:list addr_preasset) (outpl:list (bitseq (S n) * asset)%type) : 208 | (forall gamma delta, alphapre gamma = alphapre delta -> gamma = delta) -> 209 | forall j, 210 | outpl_reln txh alphapre j fulloutpl outpl -> 211 | outpl_reln txh (fun gamma => alphapre (true,gamma)) j fulloutpl (strip_bitseq_true outpl). 212 | intros Hapi j H. induction H as [j|j fulloutpl outpl alpha u H1 H2 IH|j fulloutpl outpl [[|] gamma] u H1 IH]. 213 | - simpl. apply outpl_reln_nil. 214 | - simpl. apply outpl_reln_skip. 215 | + intros gamma. apply H1. 216 | + exact IH. 217 | - simpl. 218 | apply (outpl_reln_cons txh (fun gamma => alphapre (true,gamma)) j fulloutpl (strip_bitseq_true outpl) gamma u). 219 | exact IH. 220 | - simpl. apply outpl_reln_skip. 221 | + intros gamma' H2. apply Hapi in H2. inversion H2. 222 | + exact IH. 223 | Qed. 224 | 225 | Lemma inpl_reln_nil_no_spent_lem {n:nat} (fullinpl:list addr_assetid) (alphapre:bitseq n -> addr) : 226 | inpl_reln alphapre fullinpl nil 227 | -> 228 | forall gamma:bitseq n, 229 | get_spent (alphapre gamma) fullinpl = nil. 230 | induction fullinpl as [|[alpha k] fullinpl IH]. 231 | - intros _ gamma. reflexivity. 232 | - intros H1 gamma. inversion H1. 233 | simpl. destruct (addr_eq_dec (alphapre gamma) alpha) as [D1|D1]. 234 | + exfalso. revert D1. apply H4. 235 | + now apply IH. 236 | Qed. 237 | 238 | Lemma outpl_reln_nil_no_new_assets_lem {n:nat} (fulloutpl:list addr_preasset) (txh:hashval) (alphapre:bitseq n -> addr) : 239 | forall j, 240 | outpl_reln txh alphapre j fulloutpl nil 241 | -> 242 | forall gamma:bitseq n, 243 | new_assets (alphapre gamma) fulloutpl txh j = nil. 244 | induction fulloutpl as [|[alpha [obl u]] fulloutpl IH]. 245 | - intros j _ gamma. reflexivity. 246 | - intros j H1 gamma. inversion H1. simpl. 247 | destruct (addr_eq_dec (alphapre gamma) alpha) as [D1|D1]. 248 | + exfalso. revert D1. apply H5. 249 | + now apply IH. 250 | Qed. 251 | 252 | Definition tx_inputs_valid (inpl:list addr_assetid) : Prop := 253 | inpl <> nil (*** without this, two tx can have the same hashes. ***) 254 | /\ 255 | no_dups inpl (*** inputs cannot be listed more than once. ***) 256 | . 257 | 258 | Definition tx_valid (tx:Tx) : Prop := 259 | tx_inputs_valid (tx_inputs tx). 260 | 261 | (*** Signed Transactions ***) 262 | Definition sTx : Type := prod Tx (list signat). 263 | 264 | Definition check_obligation (blockheight:nat) (h:hashval) (sl:list signat) (obl:obligation) : Prop := 265 | match obl with 266 | | (al,(m,b)) => 267 | (*** block height has been reached ***) 268 | b >= blockheight 269 | /\ 270 | (*** and for at least m of the addresses there is a signature ***) 271 | exists bl, no_dups bl 272 | /\ length bl = m 273 | /\ (forall beta, In beta bl -> In beta al 274 | /\ exists s, In s sl /\ check_signat h s beta) 275 | end. 276 | 277 | Definition tx_signatures_valid (blockheight : nat) (al:list asset) (stx:sTx) : Prop := 278 | let (tx,sl) := stx in 279 | let h := hashtx tx in 280 | forall k, (exists alpha, In (alpha,k) (tx_inputs tx)) 281 | -> exists a, In a al /\ assetid a = k /\ check_obligation blockheight h sl (assetobl a). 282 | 283 | -------------------------------------------------------------------------------- /coqcompile: -------------------------------------------------------------------------------- 1 | coqc Prelude 2 | coqc Addrs 3 | coqc CryptoHashes 4 | coqc CryptoSignatures 5 | coqc Assets 6 | coqc Transactions 7 | coqc LedgerStates 8 | coqc MTrees 9 | coqc CTrees 10 | coqc CTreeGrafting 11 | coqc Blocks 12 | -------------------------------------------------------------------------------- /lightcrypto.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/bitemyapp/ledgertheory/5f1daef98b9d8fa641a876332f901608ed9339c3/lightcrypto.pdf --------------------------------------------------------------------------------