├── .gitignore ├── BlocksWorld.v ├── EnvLemmas.v ├── LinearLogic.v ├── Makefile ├── Notes.v ├── README.txt ├── README_checkpoint.txt ├── Sokoban.v └── _CoqProject /.gitignore: -------------------------------------------------------------------------------- 1 | *~ 2 | .*.aux 3 | *.glob 4 | *.vo 5 | *.vok 6 | *.vos 7 | .lia.cache 8 | .Makefile.d 9 | .#* 10 | #* 11 | *# 12 | 2048.v -------------------------------------------------------------------------------- /BlocksWorld.v: -------------------------------------------------------------------------------- 1 | (* For the blocks world example in Power/Webster: defines primitives, defines actions of robot arms, and defines 3 sub-lemmas. 2 | 3 | Proves the main example: 4 | 5 | B 6 | A C 7 | ------- 8 | 9 | can be transformed into 10 | 11 | 12 | A 13 | B C 14 | ------- 15 | 16 | (where the second stack of 2 blocks is at a different place on the table) *) 17 | 18 | Require Import LinearLogic. 19 | Require Import EnvLemmas. 20 | Require Import Coq.Strings.String. 21 | Require Import Lia. 22 | Require Import Coq.Logic.FunctionalExtensionality. 23 | Open Scope string_scope. 24 | 25 | 26 | Definition Block : Type := string. 27 | Definition Arm : Type := string. 28 | (* Also not modeling the table *) 29 | 30 | (* Note that it's not a block datatype or arm datatype with internal state. Blocks and arms have no state. Their state is determined by the props holding for it in the environment. Could we model this differently? *) 31 | 32 | (* ----------------------- *) 33 | 34 | (* Props (describing states) *) 35 | 36 | (* Here they are just axioms (?), but they could be "more provable axioms" if each block had a coordinate and size. So to prove (on x y), you would have to show that the locations and coordinates matches up. Or, you could write a function that computed on the locs and coords and returned a boolean. *) 37 | 38 | (* Note that these are *linear logic propositions*, not *rules*, which are Coq propositions like (... --- ...) *) 39 | (* Describes blocks *) 40 | Parameter on : Block -> Block -> LinProp. 41 | Parameter table : Block -> LinProp. 42 | Parameter clear : Block -> LinProp. 43 | (* Describes robot arm *) 44 | Parameter holds : Arm -> Block -> LinProp. 45 | Parameter empty : Arm -> LinProp. 46 | 47 | (* Axioms (valid actions) *) 48 | (* Note these are rules (Coq propositions), not linear logic propositions *) 49 | 50 | (* TODO: fix sequent notation so second arg doesn't need parens *) 51 | Axiom get : forall (arm : Arm) (top bot : Block), 52 | {{ empty arm ** clear top }} 53 | |- (holds arm top ** 54 | ( (table top -o One) && (on top bot -o clear bot)) ). 55 | 56 | Axiom put : forall (top bot : Block) (arm : Arm), 57 | {{ holds arm top }} 58 | |- (empty arm ** clear top ** 59 | (table top && (clear bot -o on top bot) )). 60 | 61 | Axiom empty_eq : forall (a : Arm), eqLPC (empty a) (empty a) = true. 62 | Axiom table_eq : forall (b : Block), eqLPC (table b) (table b) = true. 63 | Axiom clear_eq : forall (b : Block), eqLPC (clear b) (clear b) = true. 64 | Axiom on_eq : forall (b c : Block), eqLPC (on b c) (on b c) = true. 65 | Axiom holds_eq : forall (a : Arm) (b : Block), eqLPC (holds a b) (holds a b) = true. 66 | 67 | (* TODO prove and move elsewhere *) 68 | Axiom times_assoc : forall A B C, (A ** (B ** C)) = ((A ** B) ** C). 69 | Axiom times_comm : forall A B, (A ** B) = (B ** A). 70 | 71 | (* ------------------------ *) 72 | 73 | (* Lemmas about actions *) 74 | 75 | (* getTable is long and messy -- first axiom I proved, so no automation *) 76 | Lemma getTable : forall (b : Block) (arm : Arm), 77 | {{ empty arm ** clear b ** table b }} 78 | |- (holds arm b). 79 | Proof. 80 | intros b arm. 81 | pose proof (get arm b b) as get. 82 | 83 | apply Times_L with (A := empty arm) (B := (clear b ** table b)). 84 | unfold inSet. 85 | unfold multiplicity. 86 | simpl. 87 | destruct (eq_neq_LinProp). lia. exfalso. apply n. 88 | unfold eqLinProp. rewrite times_assoc. simpl. rewrite empty_eq. rewrite clear_eq. rewrite table_eq. reflexivity. 89 | 90 | rewrite times_assoc. 91 | 92 | 93 | assert ( 94 | (empty arm 95 | :: (clear b) :: (table b) 96 | :: emptyBag) |- (holds arm b) 97 | -> 98 | (empty arm 99 | :: (clear b ** table b) 100 | :: {{empty arm ** clear b ** table b}} \ 101 | (empty arm ** clear b ** table b)) |- (holds arm b)). admit. apply H. clear H. 102 | 103 | (* apply Times_L with (A := clear b) (B := table b). *) 104 | (* unfold inSet. *) 105 | 106 | assert (get': 107 | ({{empty arm ** clear b}}) 108 | |- (holds arm b ** (table b -o One) && (on b b -o clear b)) -> 109 | (empty arm :: clear b :: emptyBag) 110 | |- (holds arm b ** (table b -o One) && (on b b -o clear b))). admit. 111 | apply get' in get. clear get'. 112 | 113 | apply cut with (d1 := empty arm :: clear b :: emptyBag) 114 | (d2 := table b :: emptyBag) 115 | (A := ((holds arm b) ** ((table b) -o One) && (on b b -o clear b))). 116 | unfold meq. intros. simpl. lia. 117 | 118 | apply get. 119 | 120 | (* eapply Times_L. *) 121 | 122 | assert (H: 123 | ((holds arm b) :: (table b -o One) && (on b b -o clear b) 124 | :: table b :: emptyBag) |- (holds arm b) 125 | -> 126 | ((holds arm b ** (table b -o One) && (on b b -o clear b)) 127 | :: table b :: emptyBag) |- (holds arm b)). intros. (* eapply Times_L. *) 128 | admit. 129 | 130 | apply H. clear H. 131 | 132 | (* Check With_L1. *) 133 | (* apply With_L1 with (A := (table b -o One)) (B := (on b b -o clear b)). *) 134 | (* unfold inSet. simpl. admit. *) 135 | (* TODO this will work but need to figure out how to deal with setMinus *) 136 | 137 | assert (H: 138 | (holds arm b 139 | :: (table b -o One) 140 | :: table b :: emptyBag) |- (holds arm b) 141 | -> 142 | (holds arm b 143 | :: (table b -o One) && (on b b -o clear b) 144 | :: table b :: emptyBag) |- (holds arm b)). admit. 145 | apply H. clear H. (* holds arm and empty arm? *) 146 | 147 | clear get. 148 | 149 | assert (H1: 150 | (holds arm b :: One :: emptyBag) 151 | |- (holds arm b) 152 | -> 153 | (holds arm b :: (table b -o One) :: table b :: emptyBag) 154 | |- (holds arm b)). 155 | 156 | intros. 157 | apply Impl_L with (d1 := table b :: emptyBag) (A := table b) (d2 := holds arm b :: emptyBag) (B := One). 158 | unfold inSet. simpl. remember (table b -o One) as imp. destruct (eq_neq_LinProp imp imp). 159 | lia. exfalso. apply n. unfold eqLinProp. subst. simpl. rewrite table_eq. reflexivity. 160 | (* TODO need two ltacs here!! *) 161 | 162 | unfold setMinus. simpl. unfold munion. simpl. unfold meq. intros. simpl. 163 | repeat rewrite <- plus_n_O. 164 | destruct (eq_neq_LinProp (table b -o One) a). lia. lia. 165 | 166 | constructor. unfold meq. intros. simpl. lia. 167 | 168 | admit. (* :: is "commutative" in H... may need meq and setoid rewrite *) 169 | 170 | apply H1. 171 | (* clear H H0 H1. *) 172 | 173 | apply One_L. 174 | unfold inSet. simpl. 175 | rewrite <- plus_n_O. 176 | assert (eqLinProp One One). unfold eqLinProp. simpl. reflexivity. 177 | destruct (eq_neq_LinProp One One). lia. 178 | contradiction. 179 | 180 | assert (meq ((holds arm b :: One :: emptyBag) \ One) (holds arm b :: emptyBag)). 181 | unfold setMinus. unfold munion. unfold meq. simpl. 182 | intros. repeat rewrite <- plus_n_O. destruct (eq_neq_LinProp One a). 183 | lia. lia. 184 | 185 | (* need a SETOID REWRITE with H *) 186 | assert ( 187 | (holds arm b :: emptyBag) |- (holds arm b) 188 | -> 189 | ((holds arm b :: One :: emptyBag) \ One) |- (holds arm b)). admit. 190 | apply H0. clear H0. 191 | 192 | constructor. unfold meq. intros. simpl. lia. 193 | Admitted. 194 | 195 | (* missing the fact that bot is on table, but it starts out being on the table anyway *) 196 | Lemma puton : forall (top bot : Block) (arm : Arm), 197 | {{ holds arm top ** clear bot }} 198 | |- (empty arm ** clear top ** on top bot). 199 | Proof. 200 | intros. 201 | 202 | Admitted. 203 | 204 | Lemma putTable : forall (b : Block) (arm : Arm), 205 | {{ holds arm b }} 206 | |- (empty arm ** table b ** clear b). 207 | Proof. 208 | intros. 209 | 210 | Admitted. 211 | 212 | (* -------------- *) 213 | 214 | (* Initial BlocksWorld state, goal, and proof transforming initial state into goal *) 215 | Theorem SwapAB : forall (top bot other : Block) (arm : Arm), 216 | {{ empty arm ** clear top ** on top bot ** table bot 217 | ** table other ** clear other }} 218 | |- (on bot top ** Top). 219 | (* Top is supposed to be a sink for unused predicates, since 220 | g |- T for any g. *) 221 | Proof. 222 | intros. rewrite eq_single. repeat apply unstick. apply stick. 223 | 224 | (* pick up top *) 225 | eapply cut with (d1 := {{empty arm ** clear top}}) (d2 := on top bot :: table bot :: table other :: clear other :: emptyBag). 226 | meq_clear. 227 | 228 | apply get with (top := top) (bot := bot). 229 | apply unstick. 230 | 231 | apply With_L2 with (A := (table top -o One)) (B := (on top bot -o clear bot)). 232 | inSet_clear. 233 | eqterm_clear ((table top -o One) && (on top bot -o clear bot)) n. 234 | 235 | assert ( 236 | ((on top bot -o clear bot) 237 | :: (holds arm top 238 | :: on top bot 239 | :: table bot :: table other :: clear other :: emptyBag)) 240 | == 241 | ((on top bot -o clear bot) 242 | :: (holds arm top 243 | :: (table top -o One) && (on top bot -o clear bot) 244 | :: on top bot 245 | :: table bot :: table other :: clear other :: emptyBag) \ 246 | (table top -o One) && (on top bot -o clear bot))). 247 | setMinus_clear ((table top -o One) && (on top bot -o clear bot)) a. 248 | 249 | setoid_rewrite <- H. clear H. 250 | 251 | apply Impl_L with (d1 := {{on top bot}}) (d2 := (holds arm top :: 252 | table bot :: table other :: clear other :: emptyBag)) 253 | (A := on top bot) (B := clear bot). 254 | 255 | inSet_clear. eqterm_clear (on top bot -o clear bot) n. 256 | meq_clear. destruct (eq_neq_LinProp (on top bot -o clear bot) a). 257 | simpl. lia. lia. 258 | 259 | constructor. meq_clear. 260 | 261 | (* put it on the table *) 262 | assert ((clear bot 263 | :: holds arm top :: table bot :: table other :: clear other :: emptyBag) == 264 | (holds arm top :: clear bot :: table bot :: table other :: clear other :: emptyBag)). 265 | meq_clear. 266 | 267 | setoid_rewrite H. clear H. 268 | 269 | eapply cut with (d1 := {{holds arm top}}) (d2 := clear bot :: table bot :: table other :: clear other :: emptyBag). 270 | meq_clear. 271 | 272 | eapply putTable. 273 | 274 | (* pick up bottom from table *) 275 | repeat apply unstick. 276 | assert ( 277 | (empty arm :: table bot :: clear bot 278 | :: clear top :: table top :: table other :: clear other :: emptyBag) 279 | == 280 | (empty arm :: table top :: clear top 281 | :: clear bot :: table bot :: table other :: clear other :: emptyBag)). 282 | meq_clear. 283 | setoid_rewrite <- H. clear H. 284 | 285 | apply stick. apply stick. 286 | 287 | eapply cut with (d1 := {{empty arm ** clear bot ** table bot}}) (d2 := clear top :: table top :: table other :: clear other :: emptyBag). 288 | rewrite <- times_assoc. rewrite (times_comm (table bot) (clear bot)). rewrite times_assoc. meq_clear. 289 | 290 | apply getTable. 291 | 292 | (* put it on top *) 293 | apply stick. 294 | 295 | eapply cut with (d1 := {{holds arm bot ** clear top}}) (d2 := table top :: table other :: clear other :: emptyBag). meq_clear. 296 | 297 | apply puton. 298 | 299 | (* remove unused assumptions *) 300 | (* note: no assumptions here are contradictory *) 301 | repeat apply unstick. 302 | 303 | apply Times_R with (d2 := empty arm 304 | :: clear bot :: table top :: table other :: clear other :: emptyBag) 305 | (d1 := {{on bot top}}). 306 | 307 | meq_clear. 308 | 309 | constructor. meq_clear. 310 | apply Top_R. 311 | Qed. 312 | 313 | (* Possibly a checker for validity of states: TODO *) 314 | -------------------------------------------------------------------------------- /EnvLemmas.v: -------------------------------------------------------------------------------- 1 | (* Some relations to define multiset equivalence and allow setoid rewrites. 2 | 3 | Lemmas about manipulating the environment, plus some tactics for automatically discharging environment proofs. *) 4 | 5 | Require Import LinearLogic. 6 | Require Import Setoid. 7 | Require Import Lia. 8 | Require Import String. 9 | 10 | (* Copied from BlockWorld *) 11 | Definition Block : Type := string. 12 | Definition Arm : Type := string. 13 | 14 | Parameter on : Block -> Block -> LinProp. 15 | Parameter table : Block -> LinProp. 16 | Parameter clear : Block -> LinProp. 17 | (* Describes robot arm *) 18 | Parameter holds : Arm -> Block -> LinProp. 19 | Parameter empty : Arm -> LinProp. 20 | 21 | Axiom empty_eq : forall (a : Arm), eqLPC (empty a) (empty a) = true. 22 | Axiom table_eq : forall (b : Block), eqLPC (table b) (table b) = true. 23 | Axiom clear_eq : forall (b : Block), eqLPC (clear b) (clear b) = true. 24 | Axiom on_eq : forall (b c : Block), eqLPC (on b c) (on b c) = true. 25 | Axiom holds_eq : forall (a : Arm) (b : Block), eqLPC (holds a b) (holds a b) = true. 26 | 27 | (* ---------------------------- *) 28 | 29 | Definition emptyBag := EmptyBag LinProp. 30 | 31 | (* To allow setoid rewrites on multisets *) 32 | Add Parametric Relation A : (multiset A) (@meq A) 33 | reflexivity proved by (@meq_refl A) 34 | symmetry proved by (@meq_sym A) 35 | transitivity proved by (@meq_trans A) 36 | as meq_rel. 37 | 38 | Notation "P ~= Q" := (eqLinProp P Q) (at level 60, right associativity). 39 | 40 | Lemma eqLinProp_refl : forall (A : LinProp), A ~= A. 41 | Proof. 42 | intros. unfold eqLinProp. induction A; simpl; try reflexivity; 43 | try (rewrite IHA1; rewrite IHA2; reflexivity); try assumption. 44 | symmetry. apply EqNat.beq_nat_refl. admit. 45 | Admitted. 46 | 47 | Lemma eqLinProp_sym : forall (A B : LinProp), A ~= B -> B ~= A. 48 | Proof. 49 | intros. 50 | unfold eqLinProp in *. 51 | induction A; induction B; simpl; try reflexivity; inversion H. 52 | 53 | symmetry in H1. apply EqNat.beq_nat_eq in H1. rewrite H1. reflexivity. 54 | rewrite H1. admit. 55 | 56 | Admitted. 57 | 58 | Lemma eqLinProp_trans : forall (A B C : LinProp), A ~= B -> B ~= C -> A ~= C. 59 | Proof. 60 | Admitted. 61 | 62 | Add Parametric Relation : (LinProp) (eqLinProp) 63 | reflexivity proved by (eqLinProp_refl) 64 | symmetry proved by (eqLinProp_sym) 65 | transitivity proved by (eqLinProp_trans) 66 | as eqLinProp_rel. 67 | 68 | Add Morphism LinProof with 69 | signature (@meq LinProp) ==> eqLinProp ==> (Basics.flip Basics.impl) 70 | as seq_mor. 71 | Proof. 72 | intros. 73 | Admitted. 74 | 75 | Lemma setoid_rewrite_test_sequent : forall (s: multiset LinProp), 76 | meq s emptyBag -> 77 | s |- Top. 78 | Proof. 79 | intros. 80 | setoid_rewrite H. (* works *) 81 | Admitted. 82 | 83 | (* --------------- *) 84 | 85 | (* Some automation for environments *) 86 | 87 | Tactic Notation "meq_clear" := 88 | unfold meq; 89 | intros; unfold multiplicity; simpl; 90 | try reflexivity; try lia. 91 | 92 | Tactic Notation "inSet_clear" := 93 | unfold inSet; unfold multiplicity; simpl; 94 | repeat rewrite <- plus_n_O; try lia. 95 | 96 | (* must call with ident "n" *) 97 | Tactic Notation "eqterm_clear" constr(t) ident(n) := 98 | destruct (eq_neq_LinProp t t); 99 | [ lia | 100 | exfalso; apply n; try apply eqLinProp_refl; 101 | unfold eqLinProp; simpl; 102 | try rewrite table_eq; try rewrite on_eq; 103 | try rewrite clear_eq; try rewrite holds_eq; try rewrite empty_eq; 104 | reflexivity]. 105 | 106 | Tactic Notation "setMinus_clear" constr(t) ident(a) := 107 | unfold setMinus; simpl; unfold munion; simpl; meq_clear; 108 | repeat rewrite <- plus_n_O; 109 | destruct (eq_neq_LinProp t a); 110 | lia. 111 | 112 | (* -------------- *) 113 | 114 | Lemma unstick : forall (A B C : LinProp) (e : env), 115 | (A :: B :: e) |- C -> 116 | ((A ** B) :: e) |- C. 117 | Proof. 118 | intros. 119 | apply Times_L with (A := A) (B := B). 120 | inSet_clear. destruct eq_neq_LinProp. lia. exfalso. apply n. apply eqLinProp_refl. 121 | 122 | assert ((A :: B :: ((A ** B) :: e) \ (A ** B)) == (A :: B :: e)). 123 | unfold setMinus. meq_clear. destruct (eq_neq_LinProp (A ** B) a). lia. lia. 124 | 125 | setoid_rewrite H0. assumption. 126 | Qed. 127 | 128 | (* TODO: redo to work in sequent *) 129 | Lemma unstick' : forall (A B : LinProp) (env : env), 130 | (A ** B) :: env = A :: B :: env. 131 | Proof. 132 | intros. 133 | Admitted. 134 | 135 | Lemma stick : forall (A B C : LinProp) (e : env), 136 | ((A ** B) :: e) |- C -> 137 | (A :: B :: e) |- C. 138 | Proof. 139 | intros. 140 | 141 | inversion H; subst; clear H. 142 | unfold meq in *. (* clearly true, since C must be A ** B here... *) 143 | admit. 144 | 145 | apply Impl_R. 146 | 147 | admit. 148 | Admitted. 149 | 150 | Lemma single_notation : forall (A B : LinProp), ({{A}} U {{B}}) == (A :: B :: emptyBag). 151 | Proof. meq_clear. Qed. 152 | 153 | Lemma swap : forall (A B C G : LinProp) (e : env), 154 | ((A :: (B ** C) :: e) |- G) 155 | -> 156 | (((A ** B) :: C :: e) |- G). 157 | Proof. 158 | intros. 159 | 160 | apply Times_L with (A := A) (B := B). inSet_clear. 161 | eqterm_clear (A ** B) n. 162 | 163 | assert ((A :: B :: ((A ** B) :: C :: e) \ (A ** B)) == (A :: B :: C :: e)). 164 | meq_clear. setMinus_clear (A ** B) a. rewrite H0. clear H0. 165 | 166 | inversion H; subst. admit. admit. admit. admit. 167 | (* why is this so hard to prove? it is true *) 168 | 169 | Admitted. 170 | 171 | Lemma swap' : forall (A B C G : LinProp), 172 | ((A :: {{B ** C}}) |- G) 173 | -> 174 | (((A ** B) :: {{C}}) |- G). 175 | Proof. 176 | intros. 177 | assert (forall P, {{P}} = P :: emptyBag). unfold singleton. unfold munion. 178 | unfold SingletonBag. intros. simpl. f_equal. admit. (* functional_extensionality. *) 179 | pose proof H0. 180 | specialize (H0 C). 181 | rewrite H0. 182 | apply swap. 183 | specialize (H1 (B ** C)). 184 | rewrite <- H1. 185 | apply H. 186 | Admitted. 187 | 188 | Lemma swap'' : forall (A C G : LinProp), 189 | ({{A ** C}} |- G) 190 | -> 191 | (A :: {{C}}) |- G. 192 | Proof. 193 | intros. 194 | Admitted. 195 | 196 | Lemma eq_single : forall (A : LinProp), 197 | {{A}} == (A :: emptyBag). 198 | Proof. meq_clear. Qed. 199 | -------------------------------------------------------------------------------- /LinearLogic.v: -------------------------------------------------------------------------------- 1 | (* Encode linear logic connectives and rules. (Combines Pfenning, Power/Webster, and Gwenn-Bosser's approaches.) 2 | Defines the notation for them and for manipulating the environment 3 | (multiset of linear props). *) 4 | 5 | Require Export Coq.Sets.Multiset. 6 | Require Import Coq.Arith.EqNat. 7 | Set Implicit Arguments. 8 | 9 | Definition Var : Type := nat. 10 | 11 | (* ILL connectives -- combination of those given by Pfenning + AGB's encoding *) 12 | Inductive LinProp : Type := 13 | (* Atomic *) 14 | | LProp : Var -> LinProp 15 | (* Multiplicative *) 16 | | Implies : LinProp -> LinProp -> LinProp (* -o *) 17 | | Times : LinProp -> LinProp -> LinProp (* (X) *) 18 | | One : LinProp (* Multiplicative identity TODO *) 19 | (* Additive *) 20 | | With : LinProp -> LinProp -> LinProp (* & *) 21 | | Plus : LinProp -> LinProp -> LinProp (* (+) *) 22 | | Top : LinProp (* aka True? *) 23 | | Zero : LinProp (* Additive identity TODO *) 24 | (* Exponentials *) 25 | | Bang : LinProp -> LinProp. 26 | 27 | Notation "A -o B" := (Implies A B) (at level 98, left associativity). 28 | Notation "A ** B" := (Times A B) (at level 98, left associativity). 29 | Notation "A && B" := (With A B) (at level 40, left associativity). 30 | Notation "A ++ B" := (Plus A B) (at level 60, right associativity). (* watch out *) 31 | Notation "! A" := (Bang A) (at level 98, left associativity). 32 | 33 | Definition env : Type := multiset LinProp. 34 | 35 | (* Equality of LinProps (needed for multisets) *) 36 | Fixpoint eqLPC (f1 f2 : LinProp) : bool := 37 | match f1, f2 with 38 | | One, One => true 39 | | Zero, Zero => true 40 | | Top, Top => true 41 | | LProp v1, LProp v2 => beq_nat v1 v2 42 | | Bang f1_1, Bang f2_1 => eqLPC f1_1 f2_1 43 | | Implies f1_1 f1_2, Implies f2_1 f2_2 => 44 | andb (eqLPC f1_1 f2_1) (eqLPC f1_2 f2_2) 45 | | Times f1_1 f1_2, Times f2_1 f2_2 => 46 | andb (eqLPC f1_1 f2_1) (eqLPC f1_2 f2_2) 47 | | With f1_1 f1_2, With f2_1 f2_2 => 48 | andb (eqLPC f1_1 f2_1) (eqLPC f1_2 f2_2) 49 | | _, _ => false 50 | end. 51 | 52 | Definition eqLinProp (f1 f2 : LinProp) := 53 | eqLPC f1 f2 = true. (* lift computational into propositional *) 54 | 55 | Lemma eq_neq_LinProp : forall (f1 f2 : LinProp), 56 | {eqLinProp f1 f2} + {~ eqLinProp f1 f2}. 57 | Proof. 58 | Set Printing All. 59 | intros. 60 | (* SearchAbout sumbool. *) 61 | (* Print eq_nat_decide. *) 62 | (* Print sumbool. *) 63 | destruct f1; destruct f2; try reflexivity; try auto. 64 | (* TODO: might need to use the "remember as" trick. anyway this is decidable *) 65 | Admitted. 66 | 67 | (* Things about multisets of linear props, which is how the environment is represented *) 68 | Definition singleton := SingletonBag eqLinProp eq_neq_LinProp. 69 | 70 | Definition inSet {A : Type} (m : multiset A) (x : A) : Prop := 71 | multiplicity m x > 0. 72 | 73 | Definition setMinus (m : multiset LinProp) (e : LinProp) : multiset LinProp := 74 | Bag (fun (x : LinProp) => if eq_neq_LinProp e x 75 | then multiplicity m x - 1 76 | else multiplicity m x). 77 | 78 | Notation "{{ Z }}" := (singleton Z) (at level 5, Z at level 99, right associativity). 79 | Notation "S == T" := (meq S T) (at level 1, left associativity). 80 | Notation "g1 'U' g2" := (munion g1 g2) (at level 100, right associativity). 81 | Notation "Z :: g" := (munion (singleton Z) g) (at level 60, right associativity). 82 | Notation "x ∈ S" := (inSet S x) (at level 60, right associativity). 83 | Notation "S \ x" := (setMinus S x) (at level 60, right associativity). 84 | 85 | Reserved Notation "A '|-' B" (at level 3). 86 | 87 | (* Here, (->) (Coq implication) denotes (--------) (logic "lines") *) 88 | (* convention: env name lowercase, prop name uppercase *) 89 | (* gamma = classical resources; delta = linear resources (after Pfenning) 90 | can I encode this at the type level? TODO. right now there might be problems with (!) because it doesn't distinguish *) 91 | 92 | Inductive LinProof : env -> LinProp -> Prop := 93 | | Id : forall (g : env) (A : LinProp), 94 | (g == {{A}}) -> 95 | g |- A 96 | 97 | (* Multiplicative connectives *) 98 | | Impl_R : forall (g : env) (A B : LinProp), 99 | (A :: g) |- B -> 100 | g |- (A -o B) 101 | 102 | (* basically, if you can prove the assump A, then you can have the conclusion B *) 103 | | Impl_L : forall (g d1 d2 : env) (A B C : LinProp), 104 | (A -o B) ∈ g -> 105 | (g \ (A -o B)) == (d1 U d2) -> 106 | d1 |- A -> 107 | (B :: d2) |- C -> 108 | g |- C 109 | 110 | | Times_R : forall (g d1 d2 : env) (A B : LinProp), 111 | g == (d1 U d2) -> 112 | d1 |- A -> 113 | d2 |- B -> 114 | g |- (A ** B) 115 | 116 | | Times_L : forall (g : env) (A B C : LinProp), 117 | (A ** B) ∈ g -> 118 | (A :: B :: (g \ (A ** B))) |- C -> 119 | g |- C 120 | 121 | | One_R : forall (g : env), 122 | g == (EmptyBag LinProp) -> 123 | g |- One 124 | 125 | | One_L : forall (g : env) (C : LinProp), 126 | One ∈ g -> 127 | (g \ One) |- C -> 128 | g |- C 129 | 130 | (* Additive connectives *) 131 | (* With = internal choice *) 132 | | With_R : forall (g : env) (A B : LinProp), 133 | g |- A -> 134 | g |- B -> 135 | g |- (A && B) 136 | 137 | | With_L1 : forall (g : env) (A B C : LinProp), 138 | (A && B) ∈ g -> 139 | (A :: (g \ (A && B))) |- C -> 140 | g |- C 141 | 142 | | With_L2 : forall (g : env) (A B C : LinProp), 143 | (A && B) ∈ g -> 144 | (B :: (g \ (A && B))) |- C -> 145 | g |- C 146 | 147 | | Top_R : forall (g : env), 148 | g |- Top 149 | 150 | (* (* Plus = external choice *) *) 151 | | Plus_R1 : forall (g : env) (A B : LinProp), 152 | g |- A -> 153 | g |- (A ++ B) 154 | 155 | | Plus_R2 : forall (g : env) (A B : LinProp), 156 | g |- B -> 157 | g |- (A ++ B) 158 | 159 | | Plus_L : forall (g : env) (A B C : LinProp), 160 | (A ++ B) ∈ g -> 161 | (A :: (g \ (A ++ B))) |- C -> 162 | (B :: (g \ (A ++ B))) |- C -> 163 | g |- C 164 | 165 | | Zero_L : forall (g : env) (C : LinProp), 166 | Zero ∈ g -> 167 | g |- C 168 | 169 | (* Quantifiers: included in Coq *) 170 | 171 | (* Exponentials *) 172 | (* TODO: implication is included in Coq *) 173 | 174 | (* Bang_R is a rule from Pfenning (Bang_L superseded by Bang_Replace) *) 175 | (* NOTE: the linear context has to be empty here; everything in g needs to be classical *) 176 | | Bang_R : forall (g : env) (A : LinProp), 177 | g |- A -> 178 | g |- !A 179 | 180 | | Bang_Replace : forall (g : env) (A C : LinProp), 181 | (!A) ∈ g -> 182 | (A :: (g \ (!A))) |- C -> 183 | g |- C 184 | 185 | | Bang_Replicate : forall (g : env) (A C : LinProp), 186 | (!A) ∈ g -> 187 | ((!A) :: g) |- C -> 188 | g |- C 189 | 190 | | Bang_Remove : forall (g : env) (A C : LinProp), 191 | (!A) ∈ g -> 192 | (g \ (!A)) |- C -> 193 | g |- C 194 | 195 | where "x |- y" := (LinProof x y). 196 | 197 | (* Various other ILL axioms here *) 198 | 199 | (* Linear cut: 200 | gamma is eliminated from pfenning's version *) 201 | (* Need to synthesize an A *) 202 | Axiom cut : forall (g d1 d2 : env) (A C : LinProp), 203 | g == (d1 U d2) -> 204 | d1 |- A -> 205 | (A :: d2) |- C -> 206 | g |- C. 207 | 208 | (* Factory cut -- note g' and d, not d1 and d2 *) 209 | Axiom cut_fact : forall (g g' d : env) (A C : LinProp), 210 | g == (g' U d) -> 211 | g' |- A -> 212 | ((A :: g') U d) |- C -> 213 | g |- C. 214 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | coqc LinearLogic.v EnvLemmas.v ; coqc BlocksWorld.v Sokoban.v 2 | -------------------------------------------------------------------------------- /Notes.v: -------------------------------------------------------------------------------- 1 | (* so we don't need to model classical props? what about factories? shouldn't we distinguish between linear and classical/unlimited props? ILinProp = Resource... could encode using bang? 2 | 3 | (ILL props just have type ILinProp 4 | - Name? How to represent the prop A? 5 | - it doesn't matter what value inhabits type ILinProp! e.g. we have "empty : ILinProp") *) 6 | 7 | (* ILL rules: 8 | 9 | Gamma |- A, B...N not possible -- only 10 | Gamma |- A? (where A = A (+) B... if necessary) 11 | 12 | so Gamma |- A = LinCons (or Turnstile) (Gamma/Resources : list ILinProp) (A : ILinProp) : Prop 13 | 14 | so, how can you construct/prove? a (single-consequent) sequent in ILL? 15 | well, if you want to construct/prove something of form (d1, d2, A -o B) |- C, 16 | you can do it by constructing/proving two things 17 | one of form (D1 |- A), another of form (d2 ^ B |- C)... 18 | 19 | and so on for each rule! *) 20 | 21 | (* ILL connectives: note, some are binary, some are unary, some are nullary 22 | (- Atomic proposition -- only in ILL paper -- with type Vars.t = N) 23 | 24 | * Multiplicative 25 | - Linear implication -o 26 | - Simultaneous conjunction ((X)) or ** (times) 27 | - Multiplicative truth: 1 28 | 29 | * Additive 30 | - Alternative conjunction (&) (with) 31 | - Plus (+) 32 | - True 33 | - 0 34 | 35 | * Quantifiers? 36 | - forall 37 | - exists -- TODO: actually, these are rules. also, I think they're built into Coq... 38 | unless the linear quantifier rules are different? 39 | 40 | * Exponentials? 41 | - subset 42 | - bang *) 43 | -------------------------------------------------------------------------------- /README.txt: -------------------------------------------------------------------------------- 1 | Overview 2 | 3 | I read the following papers for background: 4 | 5 | - "Working with Linear Logic in Coq" (WLLC) 6 | - chapters 1 and 3 of Pfenning's linear logic notes (PLL) 7 | - "Toward Language-Based Verification of Robot Behaviors" (LBV) 8 | - "Structural Analysis of Narratives with the Coq Proof Assistant" (SAN) 9 | - Wadler's linear logic notes 10 | 11 | Each file contains detailed comments on the contents, as well as TODOs. Here's a high-level summary of what I did. 12 | 13 | - Learned linear logic 14 | - Picked a version of linear logic to use (intuitionistic; sequent calculus with the cut rule) 15 | - Encoded its connectives, rules, and notations 16 | - This required combining environment techniques from SAN with the rules from PLL and the scenario from WLLC 17 | - Represented the environment as a multiset of linear logic props; learned how to do setoid rewrites 18 | - Formalized the blocks world example: objects, axioms, and proved* the main lemma, SwapAB 19 | - Formalized a new example, that of the box-pushing game Sokoban, and proved* a small Sokoban game 20 | 21 | * There are many admitted lemmas throughout the files. They relate almost entirely to manipulating the environment. 22 | Dealing with interactions between union, singleton, and the connective (**) required a lot of work, and it got tedious. 23 | The structure of the proofs should be evident though, since I used the linear logic rules, only admitting the env lemmas. 24 | 25 | Difficulties 26 | 27 | - The environment, setoid rewrites, and dealing with the (**) connective. 28 | - I think it would make life a lot easier to use a dedicated linear logic proof assistant instead of my shallow embedding of ILL. 29 | For example, Coq has the context, which is basically the environment e in a sequent (in "e |- g"). 30 | That takes away all the work of dealing with commutativity, associativity, etc. of things in the environment. 31 | - Inversion on linear logic proofs, or doing induction on them, didn't really work. 32 | I found it unusually hard to prove things like the following: 33 | {{ A ** B }} |- g 34 | -> 35 | A :: B :: emptyBag |- g 36 | - Short proofs were long and tedious; could use a lot more automation (to discharge env subgoals) and proof search (to apply ILL rules). 37 | 38 | Future directions 39 | 40 | - Dedicated ILL proof assistants (Twelf?). 41 | - Invariants of ILL proofs, corresponding to story or game invariants. Some examples: 42 | There exists a proof such that block C never moves. Or, for all proofs, resource X is always consumed. 43 | Or, for all proofs of a well-formed Sokoban level (that is, all traces of playing that level), 44 | the player never goes out of bounds of the walls. 45 | -------------------------------------------------------------------------------- /README_checkpoint.txt: -------------------------------------------------------------------------------- 1 | kye 2 | 5/6/15 3 | Project README 4 | 5 | Progress: 6 | 7 | - Read "Working with Linear Logic in Coq" (WLLC) 8 | - Read chapters 1 and 3 of Pfenning's linear logic notes (PLL) 9 | - Read "Toward Language-Based Verification of Robot Behaviors" (LBV) 10 | - Skimmed "Structural Analysis of Narratives with the Coq Proof Assistant" (SAN) 11 | - Skimmed Wadler's linear logic notes 12 | 13 | - Picked a version of linear logic to use (sequent calculus with bangs and the cut rule) 14 | - Encoded that version, combining environment techniques from SAN with the rules from PLL and the scenario from WLLC 15 | - Finished intuitionistic linear logic connectives, rules, notations 16 | - Started working on the blocks world: encoded states, axioms, beginning lemmas to prove, and final theorem to prove. 17 | In process of proving the first lemma. 18 | - Some problems with the environment definition. Need a multiset with setoid rewrite and element removal. 19 | Looked into multisets vs. ensembles vs. listsets vs. plain lists. Picked multisets. 20 | 21 | TODO: 22 | 23 | - Add the cut rule 24 | - Figure out how to do setoid rewrites with multisets 25 | - Possibly adjust ILL rules to make working with environments easier 26 | - Prove environment lemmas 27 | 28 | - Prove the 3 blocks world lemmas 29 | - Prove the final blocks world lemma 30 | - Add automation to this scenario 31 | (by Saturday?) 32 | 33 | - Pick a new scenario related to games or narratives 34 | - Encode it 35 | - Add automation 36 | (by Wednesday?) 37 | 38 | Difficulties: 39 | 40 | - setoid rewrite, environment manipulation 41 | - inversion on ILL rules doesn't give the intended results 42 | - still thinking about new scenarios (narratives/games) to encode. Any suggestions? 43 | One direction is something Tetris-like; another direction is something "Madame Bovary"-like 44 | -------------------------------------------------------------------------------- /Sokoban.v: -------------------------------------------------------------------------------- 1 | (* From Wikipedia: 2 | 3 | Sokoban (倉庫番 sōkoban?, warehouse keeper) is a type of transport puzzle, in which the player pushes boxes or crates around in a warehouse, trying to get them to storage locations. The puzzle is usually implemented as a video game. 4 | 5 | That is, given a freeform shape with some initial state, which contains the locations of 6 | - clear spaces (immovable, player can go on top) 7 | - one player (movable, must start on a clear space) 8 | - walls (immovable, player cannot go on top) 9 | - goals (immovable, player can go on top) 10 | - boxes (movable but change location, so player cannot go on top) 11 | 12 | The player must achieve the goal of getting each box on top of a goal, finishing with all boxes on unique goals simultaneously. 13 | 14 | The player can push boxes in one of four directions onto clear spaces. *) 15 | 16 | 17 | (* Here, we use linear logic to model the initial state, goal, objects, and allowable action. Then we solve one very easy level of Sokoban. 18 | 19 | Some notes on this file: 20 | 21 | - Why use linear logic? Because it models state. If you move forward, you consume the fact that you were on the previous space. etc. 22 | Although I mostly use the (**) connective and none of the fancier ones... 23 | - The axioms here in Coq are classical; so, the player can use them as many times as she wants. 24 | - Why do this? How's this different from writing a program in haskell? The novelty is that here, writing the rules is enough to enable you to play the game! 25 | 26 | - Sokoban is PSPACE-complete, so proof search will be slow. (It's not implemented here.) 27 | - Applying rules forward (often using cut) equates to pushing the blocks. 28 | - Applying them backward, in proof assistant style, equates to pulling the blocks. 29 | 30 | Some optimizations: 31 | 32 | - More automation, so a move rule really does a move and doesn't require the environment to be manipulated 5 times to apply it, then apply the next rule. 33 | - Pretty-printing (level -> environment parser, as well as environment -> level printer). I implemented a bit of it in the style of Coq 2048. 34 | - Sokoban only allows the player to push one box at a time. One could generalize this to chains of boxes in a line. 35 | - One could write a checker for stuck states, which do exist. *) 36 | 37 | 38 | Require Import LinearLogic. 39 | Require Import EnvLemmas. 40 | Require Import BlocksWorld. 41 | Require Import Ascii String EqNat NArith. 42 | Open Scope string_scope. 43 | 44 | Definition loc : Type := prod nat nat. 45 | 46 | Parameter clear : loc -> LinProp. 47 | Parameter player : loc -> LinProp. 48 | Parameter box : loc -> LinProp. 49 | Parameter wall : loc -> LinProp. 50 | Parameter goal : loc -> LinProp. 51 | 52 | 53 | Definition up (l : loc) := let (x,y) := l in (x, y + 1). 54 | Definition down (l : loc) := let (x,y) := l in (x, y - 1). 55 | Definition right (l : loc) := let (x,y) := l in (x + 1, y). 56 | Definition left (l : loc) := let (x,y) := l in (x - 1, y). 57 | 58 | (* The player can either move onto a clear space, or push a box one space in one direction *) 59 | Axiom move : forall (c : loc) (dir : loc -> loc), 60 | {{player (dir c)}} 61 | |- 62 | (clear (dir c) ** player c). 63 | 64 | (* Some diagrams. 65 | 66 | Forward: Push to the right 67 | Backward: Pull to the left 68 | 69 | ? p B _ ? 70 | |- 71 | ? _ p B ? 72 | 73 | Forward: Push to the left 74 | Backward: Pull to the left 75 | 76 | ? _ B p ? 77 | |- 78 | ? B p _ ? *) 79 | 80 | Axiom pushUp : forall (center : loc), 81 | {{player (down center) ** box center ** clear (up center)}} 82 | |- 83 | (clear (down center) ** player center ** box (up center)). 84 | 85 | Axiom pushDown : forall (center : loc), 86 | {{clear (down center) ** box center ** player (up center)}} 87 | |- 88 | (box (down center) ** player center ** clear (up center)). 89 | 90 | Axiom pushRight : forall (center : loc), 91 | {{player (left center) ** box center ** clear (right center)}} 92 | |- 93 | (clear (left center) ** player center ** box (right center)). 94 | 95 | Axiom pushLeft : forall (center : loc), 96 | {{clear (left center) ** box center ** player (right center)}} 97 | |- 98 | (box (left center) ** player center ** clear (right center)). 99 | 100 | (* --------------------- *) 101 | 102 | (* Solve an easy level of Sokoban (just pushes the block once to the right). 103 | Note how many environment manipulations it takes. *) 104 | 105 | Theorem level0 : 106 | {{clear (0,0)}} |- (clear (0,0)). 107 | Proof. 108 | constructor. meq_clear. 109 | Qed. 110 | 111 | (* TODO: write a level to variable parser, and variable to level *) 112 | 113 | (* Approach modeled on Coq 2048 *) 114 | Definition nl := (String (ascii_of_nat 10) "")%string. 115 | Definition test := ("hi" ++ nl ++ "hi")%string. 116 | Definition level1Str := 117 | (nl ++ 118 | "-----" ++ nl ++ 119 | "-pbg-" ++ nl ++ 120 | "-----" ++ nl ++ 121 | "")%string. 122 | 123 | Definition level1State goalLoc := 124 | wall (0,2) ** wall (1,3) ** wall (2,3) ** wall (3,3) ** wall (4,3) ** wall (5,3) ** 125 | wall (0,1) ** player (1,1) ** box (2,1) ** goal goalLoc ** clear goalLoc ** wall (4,1) ** 126 | wall (0,0) ** wall (1,0) ** wall (2,0) ** wall (3,0) ** wall (4,0) ** wall (5,0). 127 | 128 | Definition sokoban (l : Prop) (s : string) : Prop := l. 129 | 130 | Notation "[Sokoban] a" := (sokoban _ a) (at level 10). 131 | 132 | Definition toStr (s : LinProp) : string := level1Str. 133 | 134 | Tactic Notation "display" constr(s) := 135 | unfold toStr; unfold s; 136 | unfold nl; simpl; 137 | unfold ascii_of_nat; simpl; unfold ascii_of_pos; simpl. 138 | 139 | Theorem level1 : forall (goalLoc : loc) (state : LinProp), 140 | goalLoc = (3,1) -> 141 | state = level1State goalLoc -> 142 | sokoban 143 | ({{ state }} |- (box goalLoc ** Top)) 144 | (toStr state). 145 | Proof. 146 | intros. subst. display level1Str. 147 | unfold sokoban. 148 | 149 | (* TODO sokoban to sokoban ltac (re-draws board after tactic *) 150 | 151 | (* TODO automate this *) 152 | unfold level1State. 153 | 154 | (* find the relevant linprops that match the ones in pushRight; bring to front *) 155 | assert ( 156 | ((player (1, 1) ** 157 | box (2, 1) ** clear (3, 1)) ** (goal (3,1) ** 158 | wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** wall (3, 3) ** wall (4, 3) ** 159 | wall (5, 3) ** wall (0, 1) ** 160 | wall (4, 1) ** wall (0, 0) ** wall (1, 0) ** 161 | wall (2, 0) ** wall (3, 0) ** wall (4, 0) ** 162 | wall (5, 0))) 163 | = 164 | (wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** wall (3, 3) ** wall (4, 3) ** 165 | wall (5, 3) ** wall (0, 1) ** player (1, 1) ** 166 | box (2, 1) ** goal (3, 1) ** clear (3, 1) ** 167 | wall (4, 1) ** wall (0, 0) ** wall (1, 0) ** 168 | wall (2, 0) ** wall (3, 0) ** wall (4, 0) ** 169 | wall (5, 0))). (* by commutativity of (**) *) admit. 170 | 171 | rewrite <- H. clear H. 172 | 173 | (* separate into (_ ** _) :: (_ ** _) *) 174 | rewrite eq_single. 175 | apply unstick. 176 | 177 | (* then cut using the beginning and conclusion of pushRight *) 178 | eapply cut with (d1 := {{player (1, 1) ** box (2, 1) ** clear (3, 1)}}) 179 | (d2 := {{goal (3, 1) ** wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** 180 | wall (3, 3) ** wall (4, 3) ** wall (5, 3) ** 181 | wall (0, 1) ** wall (4, 1) ** wall (0, 0) ** 182 | wall (1, 0) ** wall (2, 0) ** wall (3, 0) ** 183 | wall (4, 0) ** wall (5, 0)}}). meq_clear. 184 | 185 | (* actually push *) 186 | (* rewrite times_comm. *) 187 | assert ((1,1) = left (2,1)). reflexivity. 188 | assert ((3,1) = right (2,1)). reflexivity. 189 | rewrite H. rewrite H0. clear H. clear H0. 190 | apply pushRight. 191 | 192 | (* get rid of the union, revert to ** so that the stage can be rendered *) 193 | simpl in *. 194 | 195 | repeat apply swap'. 196 | apply swap''. 197 | 198 | (* finish by cutting and applying box (3,1) *) 199 | (* need to move box out with :: *) 200 | 201 | assert ( 202 | (clear (1, 1) ** 203 | (player (2, 1) ** 204 | (box (3, 1) ** 205 | (goal (3, 1) ** wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** 206 | wall (3, 3) ** wall (4, 3) ** wall (5, 3) ** 207 | wall (0, 1) ** wall (4, 1) ** wall (0, 0) ** 208 | wall (1, 0) ** wall (2, 0) ** wall (3, 0) ** 209 | wall (4, 0) ** wall (5, 0))))) = 210 | (box (3, 1) ** 211 | (player (2, 1) ** 212 | (clear (1, 1) ** 213 | (goal (3, 1) ** wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** 214 | wall (3, 3) ** wall (4, 3) ** wall (5, 3) ** 215 | wall (0, 1) ** wall (4, 1) ** wall (0, 0) ** 216 | wall (1, 0) ** wall (2, 0) ** wall (3, 0) ** 217 | wall (4, 0) ** wall (5, 0)))))). (* by ** comm and assoc *) admit. 218 | rewrite H. clear H. 219 | 220 | rewrite eq_single. 221 | apply unstick. 222 | 223 | eapply Times_R with (d1 := {{box (3,1)}}) (d2 := {{ (player (2, 1) ** 224 | (clear (1, 1) ** 225 | (goal (3, 1) ** wall (0, 2) ** wall (1, 3) ** wall (2, 3) ** 226 | wall (3, 3) ** wall (4, 3) ** wall (5, 3) ** 227 | wall (0, 1) ** wall (4, 1) ** wall (0, 0) ** 228 | wall (1, 0) ** wall (2, 0) ** wall (3, 0) ** 229 | wall (4, 0) ** wall (5, 0)))) }}). meq_clear. 230 | 231 | constructor. meq_clear. 232 | apply Top_R. 233 | 234 | (* TODO: embed this tactic in something that manipulates the sokoban (with image) *) 235 | Admitted. 236 | 237 | 238 | Definition level2Str := 239 | " 240 | ---------- 241 | - - 242 | -p b - 243 | - - - 244 | - g - 245 | - - 246 | ---------- 247 | ". 248 | 249 | (* TODO: define this and solve it! (push block to the left, then down and to the right *) 250 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . LinearLogic 2 | 3 | LinearLogic.v 4 | EnvLemmas.v 5 | BlocksWorld.v 6 | Sokoban.v 7 | Notes.v 8 | --------------------------------------------------------------------------------