Require Export contexts_defs. Require Export lift_defs. (*#* #caption "current axioms for dropping", "base case", "untouched tail item", "dropped tail item", "updated tail item" *) (*#* #cap #alpha c in C1, e in C2, u in V, k in z, n in k, d in i, r in q *) Inductive drop: nat -> nat -> C -> C -> Prop := | drop_sort: (h,d,n:?) (drop h d (CSort n) (CSort n)) | drop_comp: (c,e:?) (drop (0) (0) c e) -> (k:?; u:?) (drop (0) (0) (CTail c k u) (CTail e k u)) | drop_drop: (k:?; h:?; c,e:?) (drop (r k h) (0) c e) -> (u:?) (drop (S h) (0) (CTail c k u) e) | drop_skip: (k:?; h,d:?; c,e:?) (drop h (r k d) c e) -> (u:?) (drop h (S d) (CTail c k (lift h (r k d) u)) (CTail e k u)). (*#* #stop file *) Hint drop : ltlc := Constructors drop. Hint discr : ltlc := Extern 4 (drop ? ? ? ?) Simpl. Section drop_gen_base. (**************************************************) Theorem drop_gen_sort: (n,h,d:?; x:?) (drop h d (CSort n) x) -> x = (CSort n). Intros until 1; InsertEq H '(CSort n); XElim H; Intros; Try Inversion H1; XAuto. Qed. Theorem drop_gen_refl: (x,e:?) (drop (0) (0) x e) -> x = e. Intros until 1; Repeat InsertEq H '(0); XElim H; Intros. (* case 1: drop_sort *) XAuto. (* case 2: drop_comp *) Rewrite H0; XAuto. (* case 3: drop_drop *) Inversion H2. (* case 4: drop_skip *) Inversion H1. Qed. Theorem drop_gen_drop: (k:?; c,x:?; u:?; h:?) (drop (S h) (0) (CTail c k u) x) -> (drop (r k h) (0) c x). Intros until 1; InsertEq H '(CTail c k u); InsertEq H '(0); InsertEq H '(S h); XElim H; Intros. (* case 1: drop_sort *) Inversion H1. (* case 2: drop_comp *) Inversion H1. (* case 3: drop_drop *) Inversion H1; Inversion H3. Rewrite <- H5; Rewrite <- H6; Rewrite <- H7; XAuto. (* case 4: drop_skip *) Inversion H2. Qed. Theorem drop_gen_skip_r: (c,x:?; u:?; h,d:?; k:?) (drop h (S d) x (CTail c k u)) -> (EX e | x = (CTail e k (lift h (r k d) u)) & (drop h (r k d) e c)). Intros; Inversion_clear H; XEAuto. Qed. Theorem drop_gen_skip_l: (c,x:?; u:?; h,d:?; k:?) (drop h (S d) (CTail c k u) x) -> (EX e v | x = (CTail e k v) & u = (lift h (r k d) v) & (drop h (r k d) c e) ). Intros; Inversion_clear H; XEAuto. Qed. End drop_gen_base. Hints Resolve drop_gen_refl : ltlc. Tactic Definition DropGenBase := Match Context With | [ H: (drop (0) (0) ?0 ?1) |- ? ] -> LApply (drop_gen_refl ?0 ?1); [ Clear H; Intros | XAuto ] | [ H: (drop ?0 ?1 (CSort ?2) ?3) |- ? ] -> LApply (drop_gen_sort ?2 ?0 ?1 ?3); [ Clear H; Intros | XAuto ] | [ H: (drop (S ?0) (0) (CTail ?1 ?2 ?3) ?4) |- ? ] -> LApply (drop_gen_drop ?2 ?1 ?4 ?3 ?0); [ Clear H; Intros | XAuto ] | [ H: (drop ?1 (S ?2) ?3 (CTail ?4 ?5 ?6)) |- ? ] -> LApply (drop_gen_skip_r ?4 ?3 ?6 ?1 ?2 ?5); [ Clear H; Intros H | XAuto ]; XElim H; Intros | [ H: (drop ?1 (S ?2) (CTail ?4 ?5 ?6) ?3) |- ? ] -> LApply (drop_gen_skip_l ?4 ?3 ?6 ?1 ?2 ?5); [ Clear H; Intros H | XAuto ]; XElim H; Intros. Section drop_props. (*****************************************************) Theorem drop_skip_bind: (h,d:?; c,e:?) (drop h d c e) -> (b:?; u:?) (drop h (S d) (CTail c (Bind b) (lift h d u)) (CTail e (Bind b) u)). Intros; Pattern 2 d; Replace d with (r (Bind b) d); XAuto. Qed. Theorem drop_refl: (c:?) (drop (0) (0) c c). XElim c; XAuto. Qed. Hints Resolve drop_refl : ltlc. Theorem drop_S: (b:?; c,e:?; u:?; h:?) (drop h (0) c (CTail e (Bind b) u)) -> (drop (S h) (0) c e). XElim c. (* case 1: CSort *) Intros; DropGenBase; Inversion H. (* case 2: CTail *) XElim h; Intros; DropGenBase. (* case 2.1: h = 0 *) Inversion H0; XAuto. (* case 2.1: h > 0 *) Apply drop_drop; RRw; XEAuto. (**) (* explicit constructor *) Qed. End drop_props. Hints Resolve drop_skip_bind drop_refl drop_S : ltlc. Tactic Definition DropS := Match Context With [ _: (drop ?1 (0) ?2 (CTail ?3 (Bind ?4) ?5)) |- ? ] -> LApply (drop_S ?4 ?2 ?3 ?5 ?1); [ Intros | XAuto ].