From 80288462f41babf77b3303c9ed758827695c3ada Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 17 May 2023 16:34:42 +0100 Subject: Add initial development files --- _CoqProject | 2 +- flake.lock | 26 + flake.nix | 65 + src/Common.v | 115 ++ src/Coqlib.v | 1312 ++++++++++++++ src/Errors.v | 194 ++ src/Hashtree.v | 599 ++++++ src/Maps.v | 1751 ++++++++++++++++++ src/Predicate.v | 1242 +++++++++++++ src/Sat.v | 580 ++++++ src/Smtpredicate.v | 5138 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 11 files changed, 11023 insertions(+), 1 deletion(-) create mode 100644 flake.lock create mode 100644 flake.nix create mode 100644 src/Common.v create mode 100644 src/Coqlib.v create mode 100644 src/Errors.v create mode 100644 src/Hashtree.v create mode 100644 src/Maps.v create mode 100644 src/Predicate.v create mode 100644 src/Sat.v create mode 100644 src/Smtpredicate.v diff --git a/_CoqProject b/_CoqProject index 50397a7..63b228c 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1 +1 @@ --R src TVSMT +-R src TVSMT \ No newline at end of file diff --git a/flake.lock b/flake.lock new file mode 100644 index 0000000..4517537 --- /dev/null +++ b/flake.lock @@ -0,0 +1,26 @@ +{ + "nodes": { + "nixpkgs": { + "locked": { + "lastModified": 1684336284, + "narHash": "sha256-JFRaAJNSa5Z3/0rcS+ot70uNBb6oP9g6fYPQdCUskcw=", + "owner": "nixos", + "repo": "nixpkgs", + "rev": "ab53f08df4fdc9cceb93b128678b5e57846360d0", + "type": "github" + }, + "original": { + "owner": "nixos", + "repo": "nixpkgs", + "type": "github" + } + }, + "root": { + "inputs": { + "nixpkgs": "nixpkgs" + } + } + }, + "root": "root", + "version": 7 +} diff --git a/flake.nix b/flake.nix new file mode 100644 index 0000000..7982b35 --- /dev/null +++ b/flake.nix @@ -0,0 +1,65 @@ +{ + description = "CompCertGSA dependencies"; + + inputs = { nixpkgs.url = "github:nixos/nixpkgs"; }; + + outputs = { self, nixpkgs }: + let + pkgs = nixpkgs.legacyPackages.x86_64-linux; + ncoq = pkgs.coq_8_13; + ncoqPackages = pkgs.coqPackages_8_13; + veriT' = pkgs.veriT.overrideAttrs (oA: { + src = pkgs.fetchurl { + url = "https://www.lri.fr/~keller/Documents-recherche/Smtcoq/veriT9f48a98.tar.gz"; + sha256 = "sha256-Pe46PxQVHWwWwx5Ei4Bl95A0otCiXZuUZ2nXuZPYnhY="; + }; + meta.broken = false; + }); + dpkgs = nixpkgs.legacyPackages.x86_64-darwin; + dncoq = dpkgs.coq_8_13; + dncoqPackages = dpkgs.coqPackages_8_13; + in { + devShell.x86_64-linux = pkgs.mkShell { + buildInputs = with pkgs; + [ ncoq + dune_2 + gcc + ncoq.ocaml + creduce +# ncoqPackages.smtcoq + zchaff + veriT' + + ncoqPackages.serapi + python3 + python3Packages.alectryon + python3Packages.sphinx_rtd_theme + ] ++ + (with ncoq.ocamlPackages; + [findlib menhir menhirLib ocamlgraph ocp-indent utop num zarith merlin]); + }; + devShell.x86_64-darwin = dpkgs.mkShell { + buildInputs = with dpkgs; + [ dncoq + dune_2 + gcc + dncoq.ocaml + dncoq.ocamlPackages.findlib + dncoq.ocamlPackages.menhir + dncoq.ocamlPackages.menhirLib + dncoq.ocamlPackages.ocamlgraph + + dncoq.ocamlPackages.ocp-indent + dncoq.ocamlPackages.utop + + ncoq.ocamlPackages.z3 + z3 + + dncoqPackages.serapi + python3 + python3Packages.alectryon + python3Packages.sphinx_rtd_theme + ]; + }; + }; +} diff --git a/src/Common.v b/src/Common.v new file mode 100644 index 0000000..2e5523b --- /dev/null +++ b/src/Common.v @@ -0,0 +1,115 @@ +Require Export Coq.Bool.Bool. +Require Export Coq.Lists.List. +Require Export Coq.Strings.String. +Require Export Coq.ZArith.ZArith. +Require Export Coq.ZArith.Znumtheory. +Require Import Coq.micromega.Lia. + +Require Export TVSMT.Coqlib. +(* Require Import compcert.lib.Integers. *) +Require Export TVSMT.Errors. + +(* This tactic due to Clement Pit-Claudel with some minor additions by JDP to + allow the result to be named: https://pit-claudel.fr/clement/MSc/#org96a1b5f *) +Inductive Learnt {A: Type} (a: A) := + | AlreadyKnown : Learnt a. + +Ltac learn_tac fact name := + lazymatch goal with + | [ H: Learnt fact |- _ ] => + fail 0 "fact" fact "has already been learnt" + | _ => let type := type of fact in + lazymatch goal with + | [ H: @Learnt type _ |- _ ] => + fail 0 "fact" fact "of type" type "was already learnt through" H + | _ => let learnt := fresh "Learn" in + pose proof (AlreadyKnown fact) as learnt; pose proof fact as name + end + end. + +Tactic Notation "learn" constr(fact) := let name := fresh "H" in learn_tac fact name. +Tactic Notation "learn" constr(fact) "as" simple_intropattern(name) := learn_tac fact name. + +Ltac unfold_rec c := unfold c; fold c. + +Ltac solve_by_inverts n := + match goal with | H : ?T |- _ => + match type of T with Prop => + inversion H; + match n with S (S (?n')) => subst; try constructor; solve_by_inverts (S n') end + end + end. + +Ltac solve_by_invert := solve_by_inverts 1. + +Ltac invert x := inversion x; subst; clear x. + +Ltac destruct_match := + match goal with + | [ |- context[match ?x with | _ => _ end ] ] => destruct x eqn:? + | [ H: context[match ?x with | _ => _ end] |- _ ] => destruct x eqn:? + end. + +Ltac auto_destruct x := destruct x eqn:?; simpl in *; try discriminate; try congruence. + +Ltac nicify_hypotheses := + repeat match goal with + | [ H : ex _ |- _ ] => invert H + | [ H : Some _ = Some _ |- _ ] => invert H + | [ H : OK _ = OK _ |- _ ] => invert H + | [ H : ?x = ?x |- _ ] => clear H + | [ H : _ /\ _ |- _ ] => invert H + | [ H : (_, _) = (_, _) |- _ ] => invert H + end. + +Ltac nicify_goals := + repeat match goal with + | [ |- _ /\ _ ] => split + | [ |- Some _ = Some _ ] => f_equal + | [ |- OK _ = OK _ ] => f_equal + | [ |- S _ = S _ ] => f_equal + end. + +Ltac kill_bools := + repeat match goal with + | [ H : _ && _ = true |- _ ] => apply andb_prop in H + | [ H : _ || _ = false |- _ ] => apply orb_false_elim in H + | [ H : negb _ = true |- _ ] => apply negb_true_iff in H + | [ H : negb _ = false |- _ ] => apply negb_false_iff in H + + | [ H : _ <=? _ = true |- _ ] => apply Z.leb_le in H + | [ H : _ <=? _ = false |- _ ] => apply Z.leb_gt in H + | [ H : _ apply Z.ltb_lt in H + | [ H : _ apply Z.ltb_ge in H + | [ H : _ >=? _ = _ |- _ ] => rewrite Z.geb_leb in H + | [ H : _ >? _ = _ |- _ ] => rewrite Z.gtb_ltb in H + + | [ H : _ =? _ = true |- _ ] => apply Z.eqb_eq in H + | [ H : _ =? _ = false |- _ ] => apply Z.eqb_neq in H + + | [ H : (_ <=? _)%positive = true |- _ ] => apply Pos.leb_le in H + | [ H : (_ <=? _)%positive = false |- _ ] => apply Pos.leb_gt in H + | [ H : (_ apply Pos.ltb_lt in H + | [ H : (_ apply Pos.ltb_ge in H + + | [ H : (_ =? _)%positive = true |- _ ] => apply Pos.eqb_eq in H + | [ H : (_ =? _)%positive = false |- _ ] => apply Pos.eqb_neq in H + end. + +Ltac substpp := + repeat match goal with + | [ H1 : ?x = Some _, H2 : ?x = Some _ |- _ ] => + let EQ := fresh "EQ" in + learn H1 as EQ; rewrite H2 in EQ; invert EQ + | _ => idtac + end. + +Ltac simplify := intros; simpl in *; + repeat progress (try nicify_hypotheses; try nicify_goals; try kill_bools; substpp); + simpl in *. + +Ltac crush := simplify; try discriminate; try congruence; try (zify; lia); + try assumption; try (solve [auto]). + +Ltac ecrush m := simplify; try discriminate; try congruence; try (zify; lia); + try assumption; try (solve [eauto with m]). diff --git a/src/Coqlib.v b/src/Coqlib.v new file mode 100644 index 0000000..99e0dbf --- /dev/null +++ b/src/Coqlib.v @@ -0,0 +1,1312 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This file collects a number of definitions and theorems that are + used throughout the development. It complements the Coq standard + library. *) + +Require Export String. +Require Export ZArith. +Require Export Znumtheory. +Require Export List. +Require Export Bool. +From Coq Require Export Lia. + +Global Set Asymmetric Patterns. + +(** * Useful tactics *) + +Ltac inv H := inversion H; clear H; subst. + +Ltac predSpec pred predspec x y := + generalize (predspec x y); case (pred x y); intro. + +Ltac caseEq name := + generalize (eq_refl name); pattern name at -1 in |- *; case name. + +Ltac destructEq name := + destruct name eqn:?. + +Ltac decEq := + match goal with + | [ |- _ = _ ] => f_equal + | [ |- (?X ?A <> ?X ?B) ] => + cut (A <> B); [intro; congruence | try discriminate] + end. + +Ltac byContradiction := exfalso. + +Lemma modusponens: forall (P Q: Prop), P -> (P -> Q) -> Q. +Proof. auto. Qed. + +Ltac exploit x := + refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _) _) + || refine (modusponens _ _ (x _ _) _) + || refine (modusponens _ _ (x _) _). + +(** * Definitions and theorems over the type [positive] *) + +Definition peq: forall (x y: positive), {x = y} + {x <> y} := Pos.eq_dec. +Global Opaque peq. + +Lemma peq_true: + forall (A: Type) (x: positive) (a b: A), (if peq x x then a else b) = a. +Proof. + intros. case (peq x x); intros. + auto. + elim n; auto. +Qed. + +Lemma peq_false: + forall (A: Type) (x y: positive) (a b: A), x <> y -> (if peq x y then a else b) = b. +Proof. + intros. case (peq x y); intros. + elim H; auto. + auto. +Qed. + +Definition Plt: positive -> positive -> Prop := Pos.lt. + +Lemma Plt_ne: + forall (x y: positive), Plt x y -> x <> y. +Proof. + unfold Plt; intros. red; intro. subst y. eelim Pos.lt_irrefl; eauto. +Qed. +Global Hint Resolve Plt_ne: coqlib. + +Lemma Plt_trans: + forall (x y z: positive), Plt x y -> Plt y z -> Plt x z. +Proof (Pos.lt_trans). + +Lemma Plt_succ: + forall (x: positive), Plt x (Pos.succ x). +Proof. + unfold Plt; intros. apply Pos.lt_succ_r. apply Pos.le_refl. +Qed. +Global Hint Resolve Plt_succ: coqlib. + +Lemma Plt_trans_succ: + forall (x y: positive), Plt x y -> Plt x (Pos.succ y). +Proof. + intros. apply Plt_trans with y. assumption. apply Plt_succ. +Qed. +Global Hint Resolve Plt_succ: coqlib. + +Lemma Plt_succ_inv: + forall (x y: positive), Plt x (Pos.succ y) -> Plt x y \/ x = y. +Proof. + unfold Plt; intros. rewrite Pos.lt_succ_r in H. + apply Pos.le_lteq; auto. +Qed. + +Definition plt (x y: positive) : {Plt x y} + {~ Plt x y}. +Proof. + unfold Plt, Pos.lt; intros. destruct (Pos.compare x y). + - right; congruence. + - left; auto. + - right; congruence. +Defined. +Global Opaque plt. + +Definition Ple: positive -> positive -> Prop := Pos.le. + +Lemma Ple_refl: forall (p: positive), Ple p p. +Proof (Pos.le_refl). + +Lemma Ple_trans: forall (p q r: positive), Ple p q -> Ple q r -> Ple p r. +Proof (Pos.le_trans). + +Lemma Plt_Ple: forall (p q: positive), Plt p q -> Ple p q. +Proof (Pos.lt_le_incl). + +Lemma Ple_succ: forall (p: positive), Ple p (Pos.succ p). +Proof. + intros. apply Plt_Ple. apply Plt_succ. +Qed. + +Lemma Plt_Ple_trans: + forall (p q r: positive), Plt p q -> Ple q r -> Plt p r. +Proof (Pos.lt_le_trans). + +Lemma Plt_strict: forall p, ~ Plt p p. +Proof (Pos.lt_irrefl). + +Global Hint Resolve Ple_refl Plt_Ple Ple_succ Plt_strict: coqlib. + +Ltac extlia := unfold Plt, Ple in *; lia. + +(** Peano recursion over positive numbers. *) + +Section POSITIVE_ITERATION. + +Lemma Plt_wf: well_founded Plt. +Proof. + apply well_founded_lt_compat with Pos.to_nat. + intros. apply nat_of_P_lt_Lt_compare_morphism. exact H. +Qed. + +Variable A: Type. +Variable v1: A. +Variable f: positive -> A -> A. + +Lemma Ppred_Plt: + forall x, x <> xH -> Plt (Pos.pred x) x. +Proof. + intros. elim (Pos.succ_pred_or x); intro. contradiction. + set (y := Pos.pred x) in *. rewrite <- H0. apply Plt_succ. +Qed. + +Let iter (x: positive) (P: forall y, Plt y x -> A) : A := + match peq x xH with + | left EQ => v1 + | right NOTEQ => f (Pos.pred x) (P (Pos.pred x) (Ppred_Plt x NOTEQ)) + end. + +Definition positive_rec : positive -> A := + Fix Plt_wf (fun _ => A) iter. + +Lemma unroll_positive_rec: + forall x, + positive_rec x = iter x (fun y _ => positive_rec y). +Proof. + unfold positive_rec. apply (Fix_eq Plt_wf (fun _ => A) iter). + intros. unfold iter. case (peq x 1); intro. auto. decEq. apply H. +Qed. + +Lemma positive_rec_base: + positive_rec 1%positive = v1. +Proof. + rewrite unroll_positive_rec. unfold iter. case (peq 1 1); intro. + auto. elim n; auto. +Qed. + +Lemma positive_rec_succ: + forall x, positive_rec (Pos.succ x) = f x (positive_rec x). +Proof. + intro. rewrite unroll_positive_rec. unfold iter. + case (peq (Pos.succ x) 1); intro. + destruct x; simpl in e; discriminate. + rewrite Pos.pred_succ. auto. +Qed. + +Lemma positive_Peano_ind: + forall (P: positive -> Prop), + P xH -> + (forall x, P x -> P (Pos.succ x)) -> + forall x, P x. +Proof. + intros. + apply (well_founded_ind Plt_wf P). + intros. + case (peq x0 xH); intro. + subst x0; auto. + elim (Pos.succ_pred_or x0); intro. contradiction. rewrite <- H2. + apply H0. apply H1. apply Ppred_Plt. auto. +Qed. + +End POSITIVE_ITERATION. + +(** * Definitions and theorems over the type [Z] *) + +Definition zeq: forall (x y: Z), {x = y} + {x <> y} := Z.eq_dec. + +Lemma zeq_true: + forall (A: Type) (x: Z) (a b: A), (if zeq x x then a else b) = a. +Proof. + intros. case (zeq x x); intros. + auto. + elim n; auto. +Qed. + +Lemma zeq_false: + forall (A: Type) (x y: Z) (a b: A), x <> y -> (if zeq x y then a else b) = b. +Proof. + intros. case (zeq x y); intros. + elim H; auto. + auto. +Qed. + +Open Scope Z_scope. + +Definition zlt: forall (x y: Z), {x < y} + {x >= y} := Z_lt_dec. + +Lemma zlt_true: + forall (A: Type) (x y: Z) (a b: A), + x < y -> (if zlt x y then a else b) = a. +Proof. + intros. case (zlt x y); intros. + auto. + extlia. +Qed. + +Lemma zlt_false: + forall (A: Type) (x y: Z) (a b: A), + x >= y -> (if zlt x y then a else b) = b. +Proof. + intros. case (zlt x y); intros. + extlia. + auto. +Qed. + +Definition zle: forall (x y: Z), {x <= y} + {x > y} := Z_le_gt_dec. + +Lemma zle_true: + forall (A: Type) (x y: Z) (a b: A), + x <= y -> (if zle x y then a else b) = a. +Proof. + intros. case (zle x y); intros. + auto. + extlia. +Qed. + +Lemma zle_false: + forall (A: Type) (x y: Z) (a b: A), + x > y -> (if zle x y then a else b) = b. +Proof. + intros. case (zle x y); intros. + extlia. + auto. +Qed. + +(** Properties of powers of two. *) + +Lemma two_power_nat_O : two_power_nat O = 1. +Proof. reflexivity. Qed. + +Lemma two_power_nat_pos : forall n : nat, two_power_nat n > 0. +Proof. + induction n. rewrite two_power_nat_O. lia. + rewrite two_power_nat_S. lia. +Qed. + +Lemma two_power_nat_two_p: + forall x, two_power_nat x = two_p (Z.of_nat x). +Proof. + induction x. auto. + rewrite two_power_nat_S. rewrite Nat2Z.inj_succ. rewrite two_p_S. lia. lia. +Qed. + +Lemma two_p_monotone: + forall x y, 0 <= x <= y -> two_p x <= two_p y. +Proof. + intros. + replace (two_p x) with (two_p x * 1) by lia. + replace y with (x + (y - x)) by lia. + rewrite two_p_is_exp; try lia. + apply Zmult_le_compat_l. + assert (two_p (y - x) > 0). apply two_p_gt_ZERO. lia. lia. + assert (two_p x > 0). apply two_p_gt_ZERO. lia. lia. +Qed. + +Lemma two_p_monotone_strict: + forall x y, 0 <= x < y -> two_p x < two_p y. +Proof. + intros. assert (two_p x <= two_p (y - 1)). apply two_p_monotone; lia. + assert (two_p (y - 1) > 0). apply two_p_gt_ZERO. lia. + replace y with (Z.succ (y - 1)) by lia. rewrite two_p_S. lia. lia. +Qed. + +Lemma two_p_strict: + forall x, x >= 0 -> x < two_p x. +Proof. + intros x0 GT. pattern x0. apply natlike_ind. + simpl. lia. + intros. rewrite two_p_S; auto. generalize (two_p_gt_ZERO x H). lia. + lia. +Qed. + +Lemma two_p_strict_2: + forall x, x >= 0 -> 2 * x - 1 < two_p x. +Proof. + intros. assert (x = 0 \/ x - 1 >= 0) by lia. destruct H0. + subst. vm_compute. auto. + replace (two_p x) with (2 * two_p (x - 1)). + generalize (two_p_strict _ H0). lia. + rewrite <- two_p_S. decEq. lia. lia. +Qed. + +(** Properties of [Zmin] and [Zmax] *) + +Lemma Zmin_spec: + forall x y, Z.min x y = if zlt x y then x else y. +Proof. + intros. case (zlt x y); unfold Z.lt, Z.ge; intro z. + unfold Z.min. rewrite z. auto. + unfold Z.min. caseEq (x ?= y); intro. + apply Z.compare_eq. auto. + contradiction. + reflexivity. +Qed. + +Lemma Zmax_spec: + forall x y, Z.max x y = if zlt y x then x else y. +Proof. + intros. case (zlt y x); unfold Z.lt, Z.ge; intro z. + unfold Z.max. rewrite <- (Zcompare_antisym y x). + rewrite z. simpl. auto. + unfold Z.max. rewrite <- (Zcompare_antisym y x). + caseEq (y ?= x); intro; simpl. + symmetry. apply Z.compare_eq. auto. + contradiction. reflexivity. +Qed. + +Lemma Zmax_bound_l: + forall x y z, x <= y -> x <= Z.max y z. +Proof. + intros. generalize (Z.le_max_l y z). lia. +Qed. +Lemma Zmax_bound_r: + forall x y z, x <= z -> x <= Z.max y z. +Proof. + intros. generalize (Z.le_max_r y z). lia. +Qed. + +(** Properties of Euclidean division and modulus. *) + +Lemma Zmod_unique: + forall x y a b, + x = a * y + b -> 0 <= b < y -> x mod y = b. +Proof. + intros. subst x. rewrite Z.add_comm. + rewrite Z_mod_plus. apply Z.mod_small. auto. lia. +Qed. + +Lemma Zdiv_unique: + forall x y a b, + x = a * y + b -> 0 <= b < y -> x / y = a. +Proof. + intros. subst x. rewrite Z.add_comm. + rewrite Z_div_plus. rewrite (Zdiv_small b y H0). lia. lia. +Qed. + +Lemma Zdiv_Zdiv: + forall a b c, + b > 0 -> c > 0 -> (a / b) / c = a / (b * c). +Proof. + intros. apply Z.div_div; lia. +Qed. + +Lemma Zdiv_interval_1: + forall lo hi a b, + lo <= 0 -> hi > 0 -> b > 0 -> + lo * b <= a < hi * b -> + lo <= a/b < hi. +Proof. + intros. + generalize (Z_div_mod_eq a b H1). generalize (Z_mod_lt a b H1). intros. + set (q := a/b) in *. set (r := a mod b) in *. + split. + assert (lo < (q + 1)). + apply Zmult_lt_reg_r with b. lia. + apply Z.le_lt_trans with a. lia. + replace ((q + 1) * b) with (b * q + b) by ring. + lia. + lia. + apply Zmult_lt_reg_r with b. lia. + replace (q * b) with (b * q) by ring. + lia. +Qed. + +Lemma Zdiv_interval_2: + forall lo hi a b, + lo <= a <= hi -> lo <= 0 -> hi >= 0 -> b > 0 -> + lo <= a/b <= hi. +Proof. + intros. + assert (lo <= a / b < hi+1). + apply Zdiv_interval_1. lia. lia. auto. + assert (lo * b <= lo * 1) by (apply Z.mul_le_mono_nonpos_l; lia). + replace (lo * 1) with lo in H3 by ring. + assert ((hi + 1) * 1 <= (hi + 1) * b) by (apply Z.mul_le_mono_nonneg_l; lia). + replace ((hi + 1) * 1) with (hi + 1) in H4 by ring. + lia. + lia. +Qed. + +Lemma Zmod_recombine: + forall x a b, + a > 0 -> b > 0 -> + x mod (a * b) = ((x/b) mod a) * b + (x mod b). +Proof. + intros. rewrite (Z.mul_comm a b). rewrite Z.rem_mul_r by lia. ring. +Qed. + +(** Properties of divisibility. *) + +Lemma Zdivide_interval: + forall a b c, + 0 < c -> 0 <= a < b -> (c | a) -> (c | b) -> 0 <= a <= b - c. +Proof. + intros. destruct H1 as [x EQ1]. destruct H2 as [y EQ2]. subst. destruct H0. + split. lia. exploit Zmult_lt_reg_r; eauto. intros. + replace (y * c - c) with ((y - 1) * c) by ring. + apply Zmult_le_compat_r; lia. +Qed. + +(** Conversion from [Z] to [nat]. *) + +Lemma Z_to_nat_neg: + forall n, n <= 0 -> Z.to_nat n = O. +Proof. + destruct n; unfold Z.le; simpl; auto. congruence. +Qed. + +Lemma Z_to_nat_max: + forall z, Z.of_nat (Z.to_nat z) = Z.max z 0. +Proof. + intros. destruct (zle 0 z). +- rewrite Z2Nat.id by auto. extlia. +- rewrite Z_to_nat_neg by lia. extlia. +Qed. + +(** Alignment: [align n amount] returns the smallest multiple of [amount] + greater than or equal to [n]. *) + +Definition align (n: Z) (amount: Z) := + ((n + amount - 1) / amount) * amount. + +Lemma align_le: forall x y, y > 0 -> x <= align x y. +Proof. + intros. unfold align. + generalize (Z_div_mod_eq (x + y - 1) y H). intro. + replace ((x + y - 1) / y * y) + with ((x + y - 1) - (x + y - 1) mod y). + generalize (Z_mod_lt (x + y - 1) y H). lia. + rewrite Z.mul_comm. lia. +Qed. + +Lemma align_divides: forall x y, y > 0 -> (y | align x y). +Proof. + intros. unfold align. apply Z.divide_factor_r. +Qed. + +(** * Definitions and theorems on the data types [option], [sum] and [list] *) + +Set Implicit Arguments. + +(** Comparing option types. *) + +Definition option_eq (A: Type) (eqA: forall (x y: A), {x=y} + {x<>y}): + forall (x y: option A), {x=y} + {x<>y}. +Proof. decide equality. Defined. +Global Opaque option_eq. + +(** Lifting a relation to an option type. *) + +Inductive option_rel (A B: Type) (R: A -> B -> Prop) : option A -> option B -> Prop := + | option_rel_none: option_rel R None None + | option_rel_some: forall x y, R x y -> option_rel R (Some x) (Some y). + +(** Mapping a function over an option type. *) + +Definition option_map (A B: Type) (f: A -> B) (x: option A) : option B := + match x with + | None => None + | Some y => Some (f y) + end. + +(** Mapping a function over a sum type. *) + +Definition sum_left_map (A B C: Type) (f: A -> B) (x: A + C) : B + C := + match x with + | inl y => inl C (f y) + | inr z => inr B z + end. + +(** Properties of [List.nth] (n-th element of a list). *) + +Global Hint Resolve in_eq in_cons: coqlib. + +Lemma nth_error_in: + forall (A: Type) (n: nat) (l: list A) (x: A), + List.nth_error l n = Some x -> In x l. +Proof. + induction n; simpl. + destruct l; intros. + discriminate. + injection H; intro; subst a. apply in_eq. + destruct l; intros. + discriminate. + apply in_cons. auto. +Qed. +Global Hint Resolve nth_error_in: coqlib. + +Lemma nth_error_nil: + forall (A: Type) (idx: nat), nth_error (@nil A) idx = None. +Proof. + induction idx; simpl; intros; reflexivity. +Qed. +Global Hint Resolve nth_error_nil: coqlib. + +(** Compute the length of a list, with result in [Z]. *) + +Fixpoint list_length_z_aux (A: Type) (l: list A) (acc: Z) {struct l}: Z := + match l with + | nil => acc + | hd :: tl => list_length_z_aux tl (Z.succ acc) + end. + +Remark list_length_z_aux_shift: + forall (A: Type) (l: list A) n m, + list_length_z_aux l n = list_length_z_aux l m + (n - m). +Proof. + induction l; intros; simpl. + lia. + replace (n - m) with (Z.succ n - Z.succ m) by lia. auto. +Qed. + +Definition list_length_z (A: Type) (l: list A) : Z := + list_length_z_aux l 0. + +Lemma list_length_z_cons: + forall (A: Type) (hd: A) (tl: list A), + list_length_z (hd :: tl) = list_length_z tl + 1. +Proof. + intros. unfold list_length_z. simpl. + rewrite (list_length_z_aux_shift tl 1 0). lia. +Qed. + +Lemma list_length_z_pos: + forall (A: Type) (l: list A), + list_length_z l >= 0. +Proof. + induction l; simpl. unfold list_length_z; simpl. lia. + rewrite list_length_z_cons. lia. +Qed. + +Lemma list_length_z_map: + forall (A B: Type) (f: A -> B) (l: list A), + list_length_z (map f l) = list_length_z l. +Proof. + induction l. reflexivity. simpl. repeat rewrite list_length_z_cons. congruence. +Qed. + +(** Extract the n-th element of a list, as [List.nth_error] does, + but the index [n] is of type [Z]. *) + +Fixpoint list_nth_z (A: Type) (l: list A) (n: Z) {struct l}: option A := + match l with + | nil => None + | hd :: tl => if zeq n 0 then Some hd else list_nth_z tl (Z.pred n) + end. + +Lemma list_nth_z_in: + forall (A: Type) (l: list A) n x, + list_nth_z l n = Some x -> In x l. +Proof. + induction l; simpl; intros. + congruence. + destruct (zeq n 0). left; congruence. right; eauto. +Qed. + +Lemma list_nth_z_map: + forall (A B: Type) (f: A -> B) (l: list A) n, + list_nth_z (List.map f l) n = option_map f (list_nth_z l n). +Proof. + induction l; simpl; intros. + auto. + destruct (zeq n 0). auto. eauto. +Qed. + +Lemma list_nth_z_range: + forall (A: Type) (l: list A) n x, + list_nth_z l n = Some x -> 0 <= n < list_length_z l. +Proof. + induction l; simpl; intros. + discriminate. + rewrite list_length_z_cons. destruct (zeq n 0). + generalize (list_length_z_pos l); lia. + exploit IHl; eauto. lia. +Qed. + +(** Properties of [List.incl] (list inclusion). *) + +Lemma incl_cons_inv: + forall (A: Type) (a: A) (b c: list A), + incl (a :: b) c -> incl b c. +Proof. + unfold incl; intros. apply H. apply in_cons. auto. +Qed. +Global Hint Resolve incl_cons_inv: coqlib. + +Lemma incl_app_inv_l: + forall (A: Type) (l1 l2 m: list A), + incl (l1 ++ l2) m -> incl l1 m. +Proof. + unfold incl; intros. apply H. apply in_or_app. left; assumption. +Qed. + +Lemma incl_app_inv_r: + forall (A: Type) (l1 l2 m: list A), + incl (l1 ++ l2) m -> incl l2 m. +Proof. + unfold incl; intros. apply H. apply in_or_app. right; assumption. +Qed. + +Global Hint Resolve incl_tl incl_refl incl_app_inv_l incl_app_inv_r: coqlib. + +Lemma incl_same_head: + forall (A: Type) (x: A) (l1 l2: list A), + incl l1 l2 -> incl (x::l1) (x::l2). +Proof. + intros; red; simpl; intros. intuition. +Qed. + +(** Properties of [List.map] (mapping a function over a list). *) + +Lemma list_map_exten: + forall (A B: Type) (f f': A -> B) (l: list A), + (forall x, In x l -> f x = f' x) -> + List.map f' l = List.map f l. +Proof. + induction l; simpl; intros. + reflexivity. + rewrite <- H. rewrite IHl. reflexivity. + intros. apply H. tauto. + tauto. +Qed. + +Lemma list_map_compose: + forall (A B C: Type) (f: A -> B) (g: B -> C) (l: list A), + List.map g (List.map f l) = List.map (fun x => g(f x)) l. +Proof. + induction l; simpl. reflexivity. rewrite IHl; reflexivity. +Qed. + +Lemma list_map_identity: + forall (A: Type) (l: list A), + List.map (fun (x:A) => x) l = l. +Proof. + induction l; simpl; congruence. +Qed. + +Lemma list_map_nth: + forall (A B: Type) (f: A -> B) (l: list A) (n: nat), + nth_error (List.map f l) n = option_map f (nth_error l n). +Proof. + induction l; simpl; intros. + repeat rewrite nth_error_nil. reflexivity. + destruct n; simpl. reflexivity. auto. +Qed. + +Lemma list_length_map: + forall (A B: Type) (f: A -> B) (l: list A), + List.length (List.map f l) = List.length l. +Proof. + induction l; simpl; congruence. +Qed. + +Lemma list_in_map_inv: + forall (A B: Type) (f: A -> B) (l: list A) (y: B), + In y (List.map f l) -> exists x:A, y = f x /\ In x l. +Proof. + induction l; simpl; intros. + contradiction. + elim H; intro. + exists a; intuition auto. + generalize (IHl y H0). intros [x [EQ IN]]. + exists x; tauto. +Qed. + +Lemma list_append_map: + forall (A B: Type) (f: A -> B) (l1 l2: list A), + List.map f (l1 ++ l2) = List.map f l1 ++ List.map f l2. +Proof. + induction l1; simpl; intros. + auto. rewrite IHl1. auto. +Qed. + +Lemma list_append_map_inv: + forall (A B: Type) (f: A -> B) (m1 m2: list B) (l: list A), + List.map f l = m1 ++ m2 -> + exists l1, exists l2, List.map f l1 = m1 /\ List.map f l2 = m2 /\ l = l1 ++ l2. +Proof. + induction m1; simpl; intros. + exists (@nil A); exists l; auto. + destruct l; simpl in H; inv H. + exploit IHm1; eauto. intros [l1 [l2 [P [Q R]]]]. subst l. + exists (a0 :: l1); exists l2; intuition. simpl; congruence. +Qed. + +(** Folding a function over a list *) + +Section LIST_FOLD. + +Variables A B: Type. +Variable f: A -> B -> B. + +(** This is exactly [List.fold_left] from Coq's standard library, + with [f] taking arguments in a different order. *) + +Fixpoint list_fold_left (accu: B) (l: list A) : B := + match l with nil => accu | x :: l' => list_fold_left (f x accu) l' end. + +(** This is exactly [List.fold_right] from Coq's standard library, + except that it runs in constant stack space. *) + +Definition list_fold_right (l: list A) (base: B) : B := + list_fold_left base (List.rev' l). + +Remark list_fold_left_app: + forall l1 l2 accu, + list_fold_left accu (l1 ++ l2) = list_fold_left (list_fold_left accu l1) l2. +Proof. + induction l1; simpl; intros. + auto. + rewrite IHl1. auto. +Qed. + +Lemma list_fold_right_eq: + forall l base, + list_fold_right l base = + match l with nil => base | x :: l' => f x (list_fold_right l' base) end. +Proof. + unfold list_fold_right; intros. + destruct l. + auto. + unfold rev'. rewrite <- ! rev_alt. simpl. + rewrite list_fold_left_app. simpl. auto. +Qed. + +Lemma list_fold_right_spec: + forall l base, list_fold_right l base = List.fold_right f base l. +Proof. + induction l; simpl; intros; rewrite list_fold_right_eq; congruence. +Qed. + +End LIST_FOLD. + +(** Properties of list membership. *) + +Lemma in_cns: + forall (A: Type) (x y: A) (l: list A), In x (y :: l) <-> y = x \/ In x l. +Proof. + intros. simpl. tauto. +Qed. + +Lemma in_app: + forall (A: Type) (x: A) (l1 l2: list A), In x (l1 ++ l2) <-> In x l1 \/ In x l2. +Proof. + intros. split; intro. apply in_app_or. auto. apply in_or_app. auto. +Qed. + +Lemma list_in_insert: + forall (A: Type) (x: A) (l1 l2: list A) (y: A), + In x (l1 ++ l2) -> In x (l1 ++ y :: l2). +Proof. + intros. apply in_or_app; simpl. elim (in_app_or _ _ _ H); intro; auto. +Qed. + +(** [list_disjoint l1 l2] holds iff [l1] and [l2] have no elements + in common. *) + +Definition list_disjoint (A: Type) (l1 l2: list A) : Prop := + forall (x y: A), In x l1 -> In y l2 -> x <> y. + +Lemma list_disjoint_cons_l: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 l2 -> ~In a l2 -> list_disjoint (a :: l1) l2. +Proof. + unfold list_disjoint; simpl; intros. destruct H1. congruence. apply H; auto. +Qed. + +Lemma list_disjoint_cons_r: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 l2 -> ~In a l1 -> list_disjoint l1 (a :: l2). +Proof. + unfold list_disjoint; simpl; intros. destruct H2. congruence. apply H; auto. +Qed. + +Lemma list_disjoint_cons_left: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint (a :: l1) l2 -> list_disjoint l1 l2. +Proof. + unfold list_disjoint; simpl; intros. apply H; tauto. +Qed. + +Lemma list_disjoint_cons_right: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 (a :: l2) -> list_disjoint l1 l2. +Proof. + unfold list_disjoint; simpl; intros. apply H; tauto. +Qed. + +Lemma list_disjoint_notin: + forall (A: Type) (l1 l2: list A) (a: A), + list_disjoint l1 l2 -> In a l1 -> ~(In a l2). +Proof. + unfold list_disjoint; intros; red; intros. + apply H with a a; auto. +Qed. + +Lemma list_disjoint_sym: + forall (A: Type) (l1 l2: list A), + list_disjoint l1 l2 -> list_disjoint l2 l1. +Proof. + unfold list_disjoint; intros. + apply not_eq_sym. apply H; auto. +Qed. + +Lemma list_disjoint_dec: + forall (A: Type) (eqA_dec: forall (x y: A), {x=y} + {x<>y}) (l1 l2: list A), + {list_disjoint l1 l2} + {~list_disjoint l1 l2}. +Proof. + induction l1; intros. + left; red; intros. elim H. + case (In_dec eqA_dec a l2); intro. + right; red; intro. apply (H a a); auto with coqlib. + case (IHl1 l2); intro. + left; red; intros. elim H; intro. + red; intro; subst a y. contradiction. + apply l; auto. + right; red; intros. elim n0. eapply list_disjoint_cons_left; eauto. +Defined. + +(** [list_equiv l1 l2] holds iff the lists [l1] and [l2] contain the same elements. *) + +Definition list_equiv (A : Type) (l1 l2: list A) : Prop := + forall x, In x l1 <-> In x l2. + +(** [list_norepet l] holds iff the list [l] contains no repetitions, + i.e. no element occurs twice. *) + +Inductive list_norepet (A: Type) : list A -> Prop := + | list_norepet_nil: + list_norepet nil + | list_norepet_cons: + forall hd tl, + ~(In hd tl) -> list_norepet tl -> list_norepet (hd :: tl). + +Lemma list_norepet_dec: + forall (A: Type) (eqA_dec: forall (x y: A), {x=y} + {x<>y}) (l: list A), + {list_norepet l} + {~list_norepet l}. +Proof. + induction l. + left; constructor. + destruct IHl. + case (In_dec eqA_dec a l); intro. + right. red; intro. inversion H. contradiction. + left. constructor; auto. + right. red; intro. inversion H. contradiction. +Defined. + +Lemma list_map_norepet: + forall (A B: Type) (f: A -> B) (l: list A), + list_norepet l -> + (forall x y, In x l -> In y l -> x <> y -> f x <> f y) -> + list_norepet (List.map f l). +Proof. + induction 1; simpl; intros. + constructor. + constructor. + red; intro. generalize (list_in_map_inv f _ _ H2). + intros [x [EQ IN]]. generalize EQ. change (f hd <> f x). + apply H1. tauto. tauto. + red; intro; subst x. contradiction. + apply IHlist_norepet. intros. apply H1. tauto. tauto. auto. +Qed. + +Remark list_norepet_append_commut: + forall (A: Type) (a b: list A), + list_norepet (a ++ b) -> list_norepet (b ++ a). +Proof. + intro A. + assert (forall (x: A) (b: list A) (a: list A), + list_norepet (a ++ b) -> ~(In x a) -> ~(In x b) -> + list_norepet (a ++ x :: b)). + induction a; simpl; intros. + constructor; auto. + inversion H. constructor. red; intro. + elim (in_app_or _ _ _ H6); intro. + elim H4. apply in_or_app. tauto. + elim H7; intro. subst a. elim H0. left. auto. + elim H4. apply in_or_app. tauto. + auto. + induction a; simpl; intros. + rewrite <- app_nil_end. auto. + inversion H0. apply H. auto. + red; intro; elim H3. apply in_or_app. tauto. + red; intro; elim H3. apply in_or_app. tauto. +Qed. + +Lemma list_norepet_app: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) <-> + list_norepet l1 /\ list_norepet l2 /\ list_disjoint l1 l2. +Proof. + induction l1; simpl; intros; split; intros. + intuition. constructor. red;simpl;auto. + tauto. + inversion H; subst. rewrite IHl1 in H3. rewrite in_app in H2. + intuition. + constructor; auto. red; intros. elim H2; intro. congruence. auto. + destruct H as [B [C D]]. inversion B; subst. + constructor. rewrite in_app. intuition. elim (D a a); auto. apply in_eq. + rewrite IHl1. intuition. red; intros. apply D; auto. apply in_cons; auto. +Qed. + +Lemma list_norepet_append: + forall (A: Type) (l1 l2: list A), + list_norepet l1 -> list_norepet l2 -> list_disjoint l1 l2 -> + list_norepet (l1 ++ l2). +Proof. + generalize list_norepet_app; firstorder. +Qed. + +Lemma list_norepet_append_right: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) -> list_norepet l2. +Proof. + generalize list_norepet_app; firstorder. +Qed. + +Lemma list_norepet_append_left: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) -> list_norepet l1. +Proof. + generalize list_norepet_app; firstorder. +Qed. + +Lemma list_norepet_rev: + forall (A: Type) (l: list A), list_norepet l -> list_norepet (List.rev l). +Proof. + induction 1; simpl. +- constructor. +- apply list_norepet_append_commut. simpl. constructor; auto. rewrite <- List.in_rev; auto. +Qed. + +(** [is_tail l1 l2] holds iff [l2] is of the form [l ++ l1] for some [l]. *) + +Inductive is_tail (A: Type): list A -> list A -> Prop := + | is_tail_refl: + forall c, is_tail c c + | is_tail_cons: + forall i c1 c2, is_tail c1 c2 -> is_tail c1 (i :: c2). + +Lemma is_tail_in: + forall (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2 -> In i c2. +Proof. + induction c2; simpl; intros. + inversion H. + inversion H. tauto. right; auto. +Qed. + +Lemma is_tail_cons_left: + forall (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2 -> is_tail c1 c2. +Proof. + induction c2; intros; inversion H. + constructor. constructor. constructor. auto. +Qed. + +Global Hint Resolve is_tail_refl is_tail_cons is_tail_in is_tail_cons_left: coqlib. + +Lemma is_tail_incl: + forall (A: Type) (l1 l2: list A), is_tail l1 l2 -> incl l1 l2. +Proof. + induction 1; eauto with coqlib. +Qed. + +Lemma is_tail_trans: + forall (A: Type) (l1 l2: list A), + is_tail l1 l2 -> forall (l3: list A), is_tail l2 l3 -> is_tail l1 l3. +Proof. + induction 1; intros. auto. apply IHis_tail. eapply is_tail_cons_left; eauto. +Qed. + +(** [list_forall2 P [x1 ... xN] [y1 ... yM]] holds iff [N = M] and + [P xi yi] holds for all [i]. *) + +Section FORALL2. + +Variable A: Type. +Variable B: Type. +Variable P: A -> B -> Prop. + +Inductive list_forall2: list A -> list B -> Prop := + | list_forall2_nil: + list_forall2 nil nil + | list_forall2_cons: + forall a1 al b1 bl, + P a1 b1 -> + list_forall2 al bl -> + list_forall2 (a1 :: al) (b1 :: bl). + +Lemma list_forall2_app: + forall a2 b2 a1 b1, + list_forall2 a1 b1 -> list_forall2 a2 b2 -> + list_forall2 (a1 ++ a2) (b1 ++ b2). +Proof. + induction 1; intros; simpl. auto. constructor; auto. +Qed. + +Lemma list_forall2_length: + forall l1 l2, + list_forall2 l1 l2 -> length l1 = length l2. +Proof. + induction 1; simpl; congruence. +Qed. + +Lemma list_forall2_in_left: + forall x1 l1 l2, + list_forall2 l1 l2 -> In x1 l1 -> exists x2, In x2 l2 /\ P x1 x2. +Proof. + induction 1; simpl; intros. contradiction. destruct H1. + subst; exists b1; auto. + exploit IHlist_forall2; eauto. intros (x2 & U & V); exists x2; auto. +Qed. + +Lemma list_forall2_in_right: + forall x2 l1 l2, + list_forall2 l1 l2 -> In x2 l2 -> exists x1, In x1 l1 /\ P x1 x2. +Proof. + induction 1; simpl; intros. contradiction. destruct H1. + subst; exists a1; auto. + exploit IHlist_forall2; eauto. intros (x1 & U & V); exists x1; auto. +Qed. + +End FORALL2. + +Lemma list_forall2_imply: + forall (A B: Type) (P1: A -> B -> Prop) (l1: list A) (l2: list B), + list_forall2 P1 l1 l2 -> + forall (P2: A -> B -> Prop), + (forall v1 v2, In v1 l1 -> In v2 l2 -> P1 v1 v2 -> P2 v1 v2) -> + list_forall2 P2 l1 l2. +Proof. + induction 1; intros. + constructor. + constructor. auto with coqlib. apply IHlist_forall2; auto. + intros. auto with coqlib. +Qed. + +(** Dropping the first N elements of a list. *) + +Fixpoint list_drop (A: Type) (n: nat) (x: list A) {struct n} : list A := + match n with + | O => x + | S n' => match x with nil => nil | hd :: tl => list_drop n' tl end + end. + +Lemma list_drop_incl: + forall (A: Type) (x: A) n (l: list A), In x (list_drop n l) -> In x l. +Proof. + induction n; simpl; intros. auto. + destruct l; auto with coqlib. +Qed. + +Lemma list_drop_norepet: + forall (A: Type) n (l: list A), list_norepet l -> list_norepet (list_drop n l). +Proof. + induction n; simpl; intros. auto. + inv H. constructor. auto. +Qed. + +Lemma list_map_drop: + forall (A B: Type) (f: A -> B) n (l: list A), + list_drop n (map f l) = map f (list_drop n l). +Proof. + induction n; simpl; intros. auto. + destruct l; simpl; auto. +Qed. + +(** * Definitions and theorems over boolean types *) + +Definition proj_sumbool {P Q: Prop} (a: {P} + {Q}) : bool := + if a then true else false. + +Coercion proj_sumbool: sumbool >-> bool. + +Lemma proj_sumbool_true: + forall (P Q: Prop) (a: {P}+{Q}), proj_sumbool a = true -> P. +Proof. + intros P Q a. destruct a; simpl. auto. congruence. +Qed. + +Lemma proj_sumbool_is_true: + forall (P: Prop) (a: {P}+{~P}), P -> proj_sumbool a = true. +Proof. + intros. unfold proj_sumbool. destruct a. auto. contradiction. +Qed. + +Ltac InvBooleans := + match goal with + | [ H: _ && _ = true |- _ ] => + destruct (andb_prop _ _ H); clear H; InvBooleans + | [ H: _ || _ = false |- _ ] => + destruct (orb_false_elim _ _ H); clear H; InvBooleans + | [ H: proj_sumbool ?x = true |- _ ] => + generalize (proj_sumbool_true _ H); clear H; intro; InvBooleans + | _ => idtac + end. + +Section DECIDABLE_EQUALITY. + +Variable A: Type. +Variable dec_eq: forall (x y: A), {x=y} + {x<>y}. +Variable B: Type. + +Lemma dec_eq_true: + forall (x: A) (ifso ifnot: B), + (if dec_eq x x then ifso else ifnot) = ifso. +Proof. + intros. destruct (dec_eq x x). auto. congruence. +Qed. + +Lemma dec_eq_false: + forall (x y: A) (ifso ifnot: B), + x <> y -> (if dec_eq x y then ifso else ifnot) = ifnot. +Proof. + intros. destruct (dec_eq x y). congruence. auto. +Qed. + +Lemma dec_eq_sym: + forall (x y: A) (ifso ifnot: B), + (if dec_eq x y then ifso else ifnot) = + (if dec_eq y x then ifso else ifnot). +Proof. + intros. destruct (dec_eq x y). + subst y. rewrite dec_eq_true. auto. + rewrite dec_eq_false; auto. +Qed. + +End DECIDABLE_EQUALITY. + +Section DECIDABLE_PREDICATE. + +Variable P: Prop. +Variable dec: {P} + {~P}. +Variable A: Type. + +Lemma pred_dec_true: + forall (a b: A), P -> (if dec then a else b) = a. +Proof. + intros. destruct dec. auto. contradiction. +Qed. + +Lemma pred_dec_false: + forall (a b: A), ~P -> (if dec then a else b) = b. +Proof. + intros. destruct dec. contradiction. auto. +Qed. + +End DECIDABLE_PREDICATE. + +(** * Well-founded orderings *) + +Require Import Relations. + +(** A non-dependent version of lexicographic ordering. *) + +Section LEX_ORDER. + +Variable A: Type. +Variable B: Type. +Variable ordA: A -> A -> Prop. +Variable ordB: B -> B -> Prop. + +Inductive lex_ord: A*B -> A*B -> Prop := + | lex_ord_left: forall a1 b1 a2 b2, + ordA a1 a2 -> lex_ord (a1,b1) (a2,b2) + | lex_ord_right: forall a b1 b2, + ordB b1 b2 -> lex_ord (a,b1) (a,b2). + +Lemma wf_lex_ord: + well_founded ordA -> well_founded ordB -> well_founded lex_ord. +Proof. + intros Awf Bwf. + assert (forall a, Acc ordA a -> forall b, Acc ordB b -> Acc lex_ord (a, b)). + induction 1. induction 1. constructor; intros. inv H3. + apply H0. auto. apply Bwf. + apply H2; auto. + red; intros. destruct a as [a b]. apply H; auto. +Qed. + +Lemma transitive_lex_ord: + transitive _ ordA -> transitive _ ordB -> transitive _ lex_ord. +Proof. + intros trA trB; red; intros. + inv H; inv H0. + left; eapply trA; eauto. + left; auto. + left; auto. + right; eapply trB; eauto. +Qed. + +End LEX_ORDER. + +(** * Nonempty lists *) + +Inductive nlist (A: Type) : Type := + | nbase: A -> nlist A + | ncons: A -> nlist A -> nlist A. + +Definition nfirst {A: Type} (l: nlist A) := + match l with nbase a => a | ncons a l' => a end. + +Fixpoint nlast {A: Type} (l: nlist A) := + match l with nbase a => a | ncons a l' => nlast l' end. + +Fixpoint nIn {A: Type} (x: A) (l: nlist A) : Prop := + match l with + | nbase a => a = x + | ncons a l => a = x \/ nIn x l + end. + +Inductive nlist_forall2 {A B: Type} (R: A -> B -> Prop) : nlist A -> nlist B -> Prop := + | nbase_forall2: forall a b, R a b -> nlist_forall2 R (nbase a) (nbase b) + | ncons_forall2: forall a l b m, R a b -> nlist_forall2 R l m -> nlist_forall2 R (ncons a l) (ncons b m). + +Lemma nlist_forall2_imply: + forall (A B: Type) (P1: A -> B -> Prop) (l1: nlist A) (l2: nlist B), + nlist_forall2 P1 l1 l2 -> + forall (P2: A -> B -> Prop), + (forall v1 v2, nIn v1 l1 -> nIn v2 l2 -> P1 v1 v2 -> P2 v1 v2) -> + nlist_forall2 P2 l1 l2. +Proof. + induction 1; simpl; intros; constructor; auto. +Qed. diff --git a/src/Errors.v b/src/Errors.v new file mode 100644 index 0000000..bf72f12 --- /dev/null +++ b/src/Errors.v @@ -0,0 +1,194 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Error reporting and the error monad. *) + +Require Import String. +Require Import Coqlib. + +Close Scope string_scope. + +Set Implicit Arguments. + +(** * Representation of error messages. *) + +(** Compile-time errors produce an error message, represented in Coq + as a list of either substrings or positive numbers encoding + a source-level identifier (see module AST). *) + +Inductive errcode: Type := + | MSG: string -> errcode + | CTX: positive -> errcode (* a top-level identifier *) + | POS: positive -> errcode. (* a positive integer, e.g. a PC *) + +Definition errmsg: Type := list errcode. + +Definition msg (s: string) : errmsg := MSG s :: nil. + +(** * The error monad *) + +(** Compilation functions that can fail have return type [res A]. + The return value is either [OK res] to indicate success, + or [Error msg] to indicate failure. *) + +Inductive res (A: Type) : Type := +| OK: A -> res A +| Error: errmsg -> res A. + +Arguments Error [A]. + +(** To automate the propagation of errors, we use a monadic style + with the following [bind] operation. *) + +Definition bind (A B: Type) (f: res A) (g: A -> res B) : res B := + match f with + | OK x => g x + | Error msg => Error msg + end. + +Definition bind2 (A B C: Type) (f: res (A * B)) (g: A -> B -> res C) : res C := + match f with + | OK (x, y) => g x y + | Error msg => Error msg + end. + +(** The [do] notation, inspired by Haskell's, keeps the code readable. *) + +Notation "'do' X <- A ; B" := (bind A (fun X => B)) + (at level 200, X ident, A at level 100, B at level 200) + : error_monad_scope. + +Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B)) + (at level 200, X ident, Y ident, A at level 100, B at level 200) + : error_monad_scope. + +Remark bind_inversion: + forall (A B: Type) (f: res A) (g: A -> res B) (y: B), + bind f g = OK y -> + exists x, f = OK x /\ g x = OK y. +Proof. + intros until y. destruct f; simpl; intros. + exists a; auto. + discriminate. +Qed. + +Remark bind2_inversion: + forall (A B C: Type) (f: res (A*B)) (g: A -> B -> res C) (z: C), + bind2 f g = OK z -> + exists x, exists y, f = OK (x, y) /\ g x y = OK z. +Proof. + intros until z. destruct f; simpl. + destruct p; simpl; intros. exists a; exists b; auto. + intros; discriminate. +Qed. + +(** Assertions *) + +Definition assertion_failed {A: Type} : res A := Error(msg "Assertion failed"). + +Notation "'assertion' A ; B" := (if A then B else assertion_failed) + (at level 200, A at level 100, B at level 200) + : error_monad_scope. + +(** This is the familiar monadic map iterator. *) + +Local Open Scope error_monad_scope. + +Fixpoint mmap (A B: Type) (f: A -> res B) (l: list A) {struct l} : res (list B) := + match l with + | nil => OK nil + | hd :: tl => do hd' <- f hd; do tl' <- mmap f tl; OK (hd' :: tl') + end. + +Remark mmap_inversion: + forall (A B: Type) (f: A -> res B) (l: list A) (l': list B), + mmap f l = OK l' -> + list_forall2 (fun x y => f x = OK y) l l'. +Proof. + induction l; simpl; intros. + inversion_clear H. constructor. + destruct (bind_inversion _ _ H) as [hd' [P Q]]. + destruct (bind_inversion _ _ Q) as [tl' [R S]]. + inversion_clear S. + constructor. auto. auto. +Qed. + +(** * Reasoning over monadic computations *) + +(** The [monadInv H] tactic below simplifies hypotheses of the form +<< + H: (do x <- a; b) = OK res +>> + By definition of the bind operation, both computations [a] and + [b] must succeed for their composition to succeed. The tactic + therefore generates the following hypotheses: + + x: ... + H1: a = OK x + H2: b x = OK res +*) + +Ltac monadInv1 H := + match type of H with + | (OK _ = OK _) => + inversion H; clear H; try subst + | (Error _ = OK _) => + discriminate + | (bind ?F ?G = OK ?X) => + let x := fresh "x" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind_inversion F G H) as [x [EQ1 EQ2]]; + clear H; + try (monadInv1 EQ2)))) + | (bind2 ?F ?G = OK ?X) => + let x1 := fresh "x" in ( + let x2 := fresh "x" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind2_inversion F G H) as [x1 [x2 [EQ1 EQ2]]]; + clear H; + try (monadInv1 EQ2))))) + | (match ?X with left _ => _ | right _ => assertion_failed end = OK _) => + destruct X; [try (monadInv1 H) | discriminate] + | (match (negb ?X) with true => _ | false => assertion_failed end = OK _) => + destruct X as [] eqn:?; simpl negb in H; [discriminate | try (monadInv1 H)] + | (match ?X with true => _ | false => assertion_failed end = OK _) => + destruct X as [] eqn:?; [try (monadInv1 H) | discriminate] + | (mmap ?F ?L = OK ?M) => + generalize (mmap_inversion F L H); intro + end. + +Ltac monadInv H := + monadInv1 H || + match type of H with + | (?F _ _ _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + end. diff --git a/src/Hashtree.v b/src/Hashtree.v new file mode 100644 index 0000000..fca7a1a --- /dev/null +++ b/src/Hashtree.v @@ -0,0 +1,599 @@ +Require Import Coq.Structures.Equalities. + +Require Import TVSMT.Maps. +Require Import TVSMT.Errors. +Require Import TVSMT.Common. + +#[local] Open Scope positive. + +#[local] Hint Resolve in_eq : core. +#[local] Hint Resolve in_cons : core. + +Declare Scope monad_scope. + +Module Type Monad. + + Parameter mon : Type -> Type. + + Parameter ret : forall (A : Type) (x : A), mon A. + Arguments ret [A]. + + Parameter bind : forall (A B : Type) (f : mon A) (g : A -> mon B), mon B. + Arguments bind [A B]. + + Parameter bind2 : forall (A B C: Type) (f: mon (A * B)) (g: A -> B -> mon C), mon C. + Arguments bind2 [A B C]. + +End Monad. + +Module MonadExtra(M : Monad). + Import M. + + Module Import MonadNotation. + + Notation "'do' X <- A ; B" := + (bind A (fun X => B)) + (at level 200, X ident, A at level 100, B at level 200) : monad_scope. + Notation "'do' ( X , Y ) <- A ; B" := + (bind2 A (fun X Y => B)) + (at level 200, X ident, Y ident, A at level 100, B at level 200) : monad_scope. + + End MonadNotation. + + #[local] Open Scope monad_scope. + + Fixpoint traverselist {A B: Type} (f: A -> mon B) (l: list A) {struct l}: mon (list B) := + match l with + | nil => ret nil + | x::xs => + do r <- f x; + do rs <- traverselist f xs; + ret (r::rs) + end. + + Fixpoint collectlist {A : Type} (f : A -> mon unit) (l : list A) {struct l} : mon unit := + match l with + | nil => ret tt + | x::xs => do y <- f x; collectlist f xs + end. + + Definition mfold_left {A B} (f: A -> B -> mon A) (l: list B) (s: mon A): mon A := + fold_left (fun a b => do a' <- a; f a' b) l s. + +End MonadExtra. + +Module Type State. + Parameter st : Type. + Parameter st_prop : st -> st -> Prop. + + Axiom st_refl : forall s, st_prop s s. + Axiom st_trans : forall s1 s2 s3, st_prop s1 s2 -> st_prop s2 s3 -> st_prop s1 s3. +End State. + +Module Statemonad(S : State) <: Monad. + + Inductive res (A: Type) (s: S.st): Type := + | Error : Errors.errmsg -> res A s + | OK : A -> forall (s' : S.st), S.st_prop s s' -> res A s. + + Arguments OK [A s]. + Arguments Error [A s]. + + Definition mon (A: Type) : Type := forall (s: S.st), res A s. + + Definition ret {A: Type} (x: A) : mon A := + fun (s : S.st) => OK x s (S.st_refl s). + + Definition bind {A B: Type} (f: mon A) (g: A -> mon B) : mon B := + fun (s : S.st) => + match f s with + | Error msg => Error msg + | OK a s' i => + match g a s' with + | Error msg => Error msg + | OK b s'' i' => OK b s'' (S.st_trans s s' s'' i i') + end + end. + + Definition bind2 {A B C: Type} (f: mon (A * B)) (g: A -> B -> mon C) : mon C := + bind f (fun xy => g (fst xy) (snd xy)). + + Definition handle_error {A: Type} (f g: mon A) : mon A := + fun (s : S.st) => + match f s with + | OK a s' i => OK a s' i + | Error _ => g s + end. + + Definition error {A: Type} (err: Errors.errmsg) : mon A := fun (s: S.st) => Error err. + + Definition get : mon S.st := fun s => OK s s (S.st_refl s). + + Definition set (s: S.st) (i: forall s', S.st_prop s' s) : mon unit := + fun s' => OK tt s (i s'). + + Definition run_mon {A: Type} (s: S.st) (m: mon A): Errors.res A := + match m s with + | OK a s' i => Errors.OK a + | Error err => Errors.Error err + end. + +End Statemonad. + +Definition max_key {A} (t: PTree.t A) := + fold_right Pos.max 1%positive (map fst (PTree.elements t)). + +Lemma max_key_correct' : + forall l hi, In hi l -> hi <= fold_right Pos.max 1 l. +Proof. + induction l; crush. + inv H. lia. + destruct (Pos.max_dec a (fold_right Pos.max 1 l)); rewrite e. + - apply Pos.max_l_iff in e. + assert (forall a b c, a <= c -> c <= b -> a <= b) by lia. + eapply H; eauto. + - apply IHl; auto. +Qed. + +Lemma max_key_correct : + forall A h_tree hi (c: A), + h_tree ! hi = Some c -> + hi <= max_key h_tree. +Proof. + unfold max_key. intros. apply PTree.elements_correct in H. + apply max_key_correct'. + eapply in_map with (f := fst) in H. auto. +Qed. + +Lemma max_not_present : + forall A k (h: PTree.t A), k > max_key h -> h ! k = None. +Proof. + intros. destruct (h ! k) eqn:?; auto. + apply max_key_correct in Heqo. lia. +Qed. + +Lemma filter_none : + forall A f l (x: A), filter f l = nil -> In x l -> f x = false. +Proof. induction l; crush; inv H0; subst; destruct_match; crush. Qed. + +Lemma filter_set : + forall A l l' f (x: A), + (In x l -> In x l') -> + In x (filter f l) -> + In x (filter f l'). +Proof. + induction l; crush. + destruct_match; crush. inv H0; crush. + apply filter_In. simplify; crush. +Qed. + +Lemma filter_cons_true : + forall A f l (a: A) l', + filter f l = a :: l' -> f a = true. +Proof. + induction l; crush. destruct (f a) eqn:?. + inv H. auto. eapply IHl; eauto. +Qed. + +Lemma PTree_set_elements : + forall A t x x' (c: A), + In x (PTree.elements t) -> + x' <> (fst x) -> + In x (PTree.elements (PTree.set x' c t)). +Proof. + intros. destruct x. + eapply PTree.elements_correct. simplify. + rewrite PTree.gso; auto. apply PTree.elements_complete in H. auto. +Qed. + +Lemma filter_set2 : + forall A x y z (h: PTree.t A), + In z (PTree.elements (PTree.set x y h)) -> + In z (PTree.elements h) \/ fst z = x. +Proof. + intros. destruct z. + destruct (Pos.eq_dec p x); subst. + tauto. + left. apply PTree.elements_correct. apply PTree.elements_complete in H. + rewrite PTree.gso in H; auto. +Qed. + +Lemma in_filter : forall A f l (x: A), In x (filter f l) -> In x l. +Proof. induction l; crush. destruct_match; crush. inv H; crush. Qed. + +Lemma filter_norepet: + forall A f (l: list A), + list_norepet l -> + list_norepet (filter f l). +Proof. + induction l; crush. + inv H. destruct (f a). + constructor. unfold not in *; intros. apply H2. + eapply in_filter; eauto. + apply IHl; auto. + apply IHl; auto. +Qed. + +Lemma filter_norepet2: + forall A B g (l: list (A * B)), + list_norepet (map fst l) -> + list_norepet (map fst (filter g l)). +Proof. + induction l; crush. + inv H. destruct (g a) eqn:?. + simplify. constructor. unfold not in *. intros. + eapply H2. + apply list_in_map_inv in H. simplify; subst. + rewrite H. + apply filter_In in H1. simplify. + apply in_map. eauto. + eapply IHl. eauto. + eapply IHl. eauto. +Qed. + +Module Type Hashable := UsualDecidableType. + +Module HashTree(Import H: Hashable). + + Definition hash := positive. + Definition hash_tree := PTree.t t. + + Definition empty := PTree.empty t. + + Definition find_tree (el: t) (h: hash_tree) : option hash := + match filter (fun x => if eq_dec el (snd x) then true else false) (PTree.elements h) with + | (p, _) :: nil => Some p + | _ => None + end. + + Definition hash_value (max: hash) (e: t) (h: hash_tree): hash * hash_tree := + match find_tree e h with + | Some p => (p, h) + | None => + let nkey := Pos.max max (max_key h) + 1 in + (nkey, PTree.set nkey e h) + end. + + Definition hash_value2 (max: hash) (e: t * t) (h: hash_tree): (hash * hash) * hash_tree := + let (e1, e2) := e in + let (v1, h1) := hash_value max e1 h in + let (v2, h2) := hash_value max e2 h1 in + ((v1, v2), h2). + + Definition hash_value3 (max: hash) (e: t * t * t) (h: hash_tree): (hash * hash * hash) * hash_tree := + let '(e1, e2, e3) := e in + let (v1, h1) := hash_value max e1 h in + let (v2, h2) := hash_value max e2 h1 in + let (v3, h3) := hash_value max e3 h2 in + ((v1, v2, v3), h3). + + Definition hash_value4 (max: hash) (e: t * t * t * t) (h: hash_tree): (hash * hash * hash * hash) * hash_tree := + let '(e1, e2, e3, e4) := e in + let (v1, h1) := hash_value max e1 h in + let (v2, h2) := hash_value max e2 h1 in + let (v3, h3) := hash_value max e3 h2 in + let (v4, h4) := hash_value max e4 h3 in + ((v1, v2, v3, v4), h4). + + Definition hash_list (max: hash) (e: list t) (h: hash_tree): list hash * hash_tree := + fold_left (fun s e => + let (i, h') := hash_value max e (snd s) in + (fst s ++ i::nil, h') + ) e (nil, h). + + Definition wf_hash_table h_tree := + (forall x c, h_tree ! x = Some c -> find_tree c h_tree = Some x). + + Lemma find_tree_correct : + forall c h_tree p, + find_tree c h_tree = Some p -> + h_tree ! p = Some c. + Proof. + intros. + unfold find_tree in H. destruct_match; crush. + destruct_match; simplify. + destruct_match; crush. + assert (In (p, t0) (filter + (fun x : hash * t => + if eq_dec c (snd x) then true else false) (PTree.elements h_tree))). + { setoid_rewrite Heql. constructor; auto. } + apply filter_In in H. simplify. destruct_match; crush. subst. + apply PTree.elements_complete; auto. + Qed. + + Lemma find_tree_unique : + forall c h_tree p p', + find_tree c h_tree = Some p -> + h_tree ! p' = Some c -> + p = p'. + Proof. + intros. + unfold find_tree in H. + repeat (destruct_match; crush; []). + assert (In (p, t0) (filter + (fun x : hash * t => + if eq_dec c (snd x) then true else false) (PTree.elements h_tree))). + { setoid_rewrite Heql. constructor; auto. } + apply filter_In in H. simplify. + destruct (Pos.eq_dec p p'); auto. + exfalso. + destruct_match; subst; crush. + assert (In (p', t0) (PTree.elements h_tree) /\ (fun x : hash * t => + if eq_dec t0 (snd x) then true else false) (p', t0) = true). + { split. apply PTree.elements_correct. auto. setoid_rewrite Heqs. auto. } + apply filter_In in H. setoid_rewrite Heql in H. inv H. simplify. crush. crush. + Qed. + + Lemma hash_no_element' : + forall c h_tree, + find_tree c h_tree = None -> + wf_hash_table h_tree -> + ~ forall x, h_tree ! x = Some c. + Proof. + unfold not, wf_hash_table; intros. + specialize (H1 1). eapply H0 in H1. crush. + Qed. + + Lemma hash_no_element : + forall c h_tree, + find_tree c h_tree = None -> + wf_hash_table h_tree -> + ~ exists x, h_tree ! x = Some c. + Proof. + unfold not, wf_hash_table; intros. + simplify. apply H0 in H2. rewrite H in H2. crush. + Qed. + + Lemma wf_hash_table_set_gso' : + forall h x p0 c', + filter + (fun x : hash * t => + if eq_dec c' (snd x) then true else false) (PTree.elements h) = + (x, p0) :: nil -> + h ! x = Some p0 /\ p0 = c'. + Proof. + intros. + match goal with + | H: filter ?f ?el = ?x::?xs |- _ => + assert (In x (filter f el)) by (rewrite H; crush) + end. + apply filter_In in H0. simplify. destruct_match; subst; crush. + apply PTree.elements_complete; auto. + destruct_match; crush. + Qed. + + Lemma wf_hash_table_set_gso : + forall x x' c' c h, + x <> x' -> + wf_hash_table h -> + find_tree c' h = Some x -> + find_tree c h = None -> + find_tree c' (PTree.set x' c h) = Some x. + Proof. + intros. pose proof H1 as X. unfold find_tree in H1. + destruct_match; crush. + destruct p. destruct l; crush. + apply wf_hash_table_set_gso' in Heql. simplify. + pose proof H2 as Z. apply hash_no_element in H2; auto. + destruct (eq_dec c c'); subst. + { exfalso. eapply H2. econstructor; eauto. } + unfold wf_hash_table in H0. + assert (In (x', c) (PTree.elements (PTree.set x' c h))). + { apply PTree.elements_correct. rewrite PTree.gss; auto. } + assert (In (x, c') (PTree.elements h)). + { apply PTree.elements_correct; auto. } + assert (In (x, c') (PTree.elements (PTree.set x' c h))). + { apply PTree.elements_correct. rewrite PTree.gso; auto. } + pose proof X as Y. + unfold find_tree in X. repeat (destruct_match; crush; []). + match goal with + | H: filter ?f ?el = ?x::?xs |- _ => + assert (In x (filter f el)) by (rewrite H; crush) + end. + apply filter_In in H6. simplify. destruct_match; crush; subst. + unfold find_tree. repeat (destruct_match; crush). + { eapply filter_none in Heql0. + 2: { apply PTree.elements_correct. rewrite PTree.gso; eauto. } + destruct_match; crush. } + + { subst. + repeat match goal with + | H: filter ?f ?el = ?x::?xs |- _ => + learn H; assert (In x (filter f el)) by (rewrite H; crush) + end. + eapply filter_set in H10. rewrite Heql0 in H10. inv H10. simplify. auto. + inv H11. auto. } + + { exfalso. subst. + repeat match goal with + | H: filter ?f ?el = ?x::?xs |- _ => + learn H; assert (In x (filter f el)) by (rewrite H; crush) + end. + + pose proof H8 as X2. destruct p1. + pose proof X2 as X4. + apply in_filter in X2. apply PTree.elements_complete in X2. + assert (In (p, t2) (filter + (fun x : positive * t => if eq_dec t0 (snd x) then true else false) + (PTree.elements (PTree.set x' c h)))) by (rewrite H6; eauto). + pose proof H11 as X3. + apply in_filter in H11. apply PTree.elements_complete in H11. + destruct (peq p0 p); subst. + { + assert (list_norepet (map fst (filter + (fun x : positive * t => if eq_dec t0 (snd x) then true else false) + (PTree.elements (PTree.set x' c h))))). + { eapply filter_norepet2. eapply PTree.elements_keys_norepet. } + rewrite Heql0 in H12. simplify. inv H12. eapply H15. apply in_eq. + } + { apply filter_In in X4. simplify. destruct_match; crush; subst. + apply filter_In in X3. simplify. destruct_match; crush; subst. + destruct (peq p x'); subst. + { rewrite PTree.gss in H11; crush. } + { destruct (peq p0 x'); subst. + { rewrite PTree.gss in X2; crush. } + { rewrite PTree.gso in X2 by auto. + rewrite PTree.gso in H11 by auto. + assert (p = p0) by (eapply find_tree_unique; eauto). + crush. } } } } + Qed. + + Lemma wf_hash_table_set : + forall h_tree c v (GT: v > max_key h_tree), + find_tree c h_tree = None -> + wf_hash_table h_tree -> + wf_hash_table (PTree.set v c h_tree). + Proof. + unfold wf_hash_table; simplify. + destruct (peq v x); subst. + pose proof (hash_no_element c h_tree H H0). + rewrite PTree.gss in H1. simplify. + unfold find_tree. + assert (In (x, c0) (PTree.elements (PTree.set x c0 h_tree)) + /\ (fun x : positive * t => if eq_dec c0 (snd x) then true else false) + (x, c0) = true). + { simplify. apply PTree.elements_correct. rewrite PTree.gss. auto. + destruct (eq_dec c0 c0); crush. } + destruct_match. + apply filter_In in H1. rewrite Heql in H1. crush. + apply filter_In in H1. repeat (destruct_match; crush; []). subst. + destruct l. simplify. rewrite Heql in H1. inv H1. inv H3. auto. + crush. + + exfalso. apply H2. destruct p. + pose proof Heql as X. apply filter_cons_true in X. destruct_match; crush; subst. + assert (In (p0, t0) (filter + (fun x : positive * t => if eq_dec t0 (snd x) then true else false) + (PTree.elements (PTree.set x t0 h_tree)))) by (rewrite Heql; eauto). + assert (In (p, t1) (filter + (fun x : positive * t => if eq_dec t0 (snd x) then true else false) + (PTree.elements (PTree.set x t0 h_tree)))) by (rewrite Heql; eauto). + apply filter_In in H4. simplify. destruct_match; crush; subst. + apply in_filter in H3. apply PTree.elements_complete in H5. apply PTree.elements_complete in H3. + assert (list_norepet (map fst (filter + (fun x : positive * t => if eq_dec t1 (snd x) then true else false) + (PTree.elements (PTree.set x t1 h_tree))))). + { eapply filter_norepet2. eapply PTree.elements_keys_norepet. } + rewrite Heql in H4. simplify. + destruct (peq p0 p); subst. + { inv H4. exfalso. eapply H8. eauto. } + destruct (peq x p); subst. + rewrite PTree.gso in H3; auto. econstructor; eauto. + rewrite PTree.gso in H5; auto. econstructor; eauto. + + rewrite PTree.gso in H1; auto. + destruct (eq_dec c c0); subst. + { apply H0 in H1. rewrite H in H1. discriminate. } + apply H0 in H1. + apply wf_hash_table_set_gso; eauto. + Qed. + + Lemma wf_hash_table_distr : + forall m p h_tree h h_tree', + hash_value m p h_tree = (h, h_tree') -> + wf_hash_table h_tree -> + wf_hash_table h_tree'. + Proof. + unfold hash_value; simplify. + destruct_match. + - inv H; auto. + - inv H. apply wf_hash_table_set; try lia; auto. + Qed. + + Lemma wf_hash_table_eq : + forall h_tree a b c, + wf_hash_table h_tree -> + h_tree ! a = Some c -> + h_tree ! b = Some c -> + a = b. + Proof. + unfold wf_hash_table; intros; apply H in H0; eapply find_tree_unique; eauto. + Qed. + + Lemma hash_constant : + forall p h h_tree hi c h_tree' m, + h_tree ! hi = Some c -> + hash_value m p h_tree = (h, h_tree') -> + h_tree' ! hi = Some c. + Proof. + intros. unfold hash_value in H0. destruct_match. + inv H0. eauto. + inv H0. + pose proof H. apply max_key_correct in H0. + rewrite PTree.gso; solve [eauto | lia]. + Qed. + + Lemma find_tree_Some : + forall el h v, + find_tree el h = Some v -> + h ! v = Some el. + Proof. + intros. unfold find_tree in *. + destruct_match; crush. destruct p. + destruct_match; crush. + match goal with + | H: filter ?f ?el = ?x::?xs |- _ => + assert (In x (filter f el)) by (rewrite H; crush) + end. + apply PTree.elements_complete. + apply filter_In in H. inv H. + destruct_match; crush. + Qed. + + Lemma hash_present_eq : + forall m e1 e2 p1 h h', + hash_value m e2 h = (p1, h') -> + h ! p1 = Some e1 -> e1 = e2. + Proof. + intros. unfold hash_value in *. destruct_match. + - inv H. apply find_tree_Some in Heqo. + rewrite Heqo in H0. inv H0. auto. + - inv H. assert (h ! (Pos.max m (max_key h) + 1) = None) + by (apply max_not_present; lia). crush. + Qed. + + Module HashState <: State. + Definition st := hash_tree. + Definition st_prop (h1 h2: hash_tree): Prop := + forall x y, h1 ! x = Some y -> h2 ! x = Some y. + + Lemma st_refl : + forall s, st_prop s s. + Proof. unfold st_prop; auto. Qed. + + Lemma st_trans : + forall s1 s2 s3, + st_prop s1 s2 -> st_prop s2 s3 -> st_prop s1 s3. + Proof. + unfold st_prop; intros; eauto. + Qed. + + #[global] Instance HashStatePreorder : PreOrder st_prop := + { PreOrder_Reflexive := st_refl; + PreOrder_Transitive := st_trans; + }. + End HashState. + + Module HashMonad := Statemonad(HashState). + +End HashTree. + +Definition gt_1 {A} h := + forall x (y: A), h ! x = Some y -> 1 < x. + +Module HashTreeProperties (Import H: Hashable). + + Module Import HT := HashTree(H). + + Lemma hash_value_gt : + forall max v h, + gt_1 h -> + gt_1 (snd (hash_value max v h)). + Proof. + unfold gt_1, hash_value; intros. + destruct_match; eauto. + destruct (peq (Pos.max max (max_key h) + 1) x); [subst;lia|]. + cbn [snd] in *. rewrite PTree.gso in H0; eauto. + Qed. + +End HashTreeProperties. diff --git a/src/Maps.v b/src/Maps.v new file mode 100644 index 0000000..6bc6e14 --- /dev/null +++ b/src/Maps.v @@ -0,0 +1,1751 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Applicative finite maps are the main data structure used in this + project. A finite map associates data to keys. The two main operations + are [set k d m], which returns a map identical to [m] except that [d] + is associated to [k], and [get k m] which returns the data associated + to key [k] in map [m]. In this library, we distinguish two kinds of maps: +- Trees: the [get] operation returns an option type, either [None] + if no data is associated to the key, or [Some d] otherwise. +- Maps: the [get] operation always returns a data. If no data was explicitly + associated with the key, a default data provided at map initialization time + is returned. + + In this library, we provide efficient implementations of trees and + maps whose keys range over the type [positive] of binary positive + integers or any type that can be injected into [positive]. The + implementation is based on radix-2 search trees (uncompressed + Patricia trees) and guarantees logarithmic-time operations. An + inefficient implementation of maps as functions is also provided. +*) + +Require Import Equivalence EquivDec. +Require Import Coqlib. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +Set Implicit Arguments. + +(** * The abstract signatures of trees *) + +Module Type TREE. + Parameter elt: Type. + Parameter elt_eq: forall (a b: elt), {a = b} + {a <> b}. + Parameter t: Type -> Type. + Parameter empty: forall (A: Type), t A. + Parameter get: forall (A: Type), elt -> t A -> option A. + Parameter set: forall (A: Type), elt -> A -> t A -> t A. + Parameter remove: forall (A: Type), elt -> t A -> t A. + + (** The ``good variables'' properties for trees, expressing + commutations between [get], [set] and [remove]. *) + Axiom gempty: + forall (A: Type) (i: elt), get i (empty A) = None. + Axiom gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = Some x. + Axiom gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Axiom gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then Some x else get i m. + (* We could implement the following, but it's not needed for the moment. + Hypothesis gsident: + forall (A: Type) (i: elt) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + Hypothesis grident: + forall (A: Type) (i: elt) (m: t A) (v: A), + get i m = None -> remove i m = m. + *) + Axiom grs: + forall (A: Type) (i: elt) (m: t A), get i (remove i m) = None. + Axiom gro: + forall (A: Type) (i j: elt) (m: t A), + i <> j -> get i (remove j m) = get i m. + Axiom grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + + (** Extensional equality between trees. *) + Parameter beq: forall (A: Type), (A -> A -> bool) -> t A -> t A -> bool. + Axiom beq_correct: + forall (A: Type) (eqA: A -> A -> bool) (t1 t2: t A), + beq eqA t1 t2 = true <-> + (forall (x: elt), + match get x t1, get x t2 with + | None, None => True + | Some y1, Some y2 => eqA y1 y2 = true + | _, _ => False + end). + + (** Applying a function to all data of a tree. *) + Parameter map: + forall (A B: Type), (elt -> A -> B) -> t A -> t B. + Axiom gmap: + forall (A B: Type) (f: elt -> A -> B) (i: elt) (m: t A), + get i (map f m) = option_map (f i) (get i m). + + (** Same as [map], but the function does not receive the [elt] argument. *) + Parameter map1: + forall (A B: Type), (A -> B) -> t A -> t B. + Axiom gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + + (** Applying a function pairwise to all data of two trees. *) + Parameter combine: + forall (A B C: Type), (option A -> option B -> option C) -> t A -> t B -> t C. + Axiom gcombine: + forall (A B C: Type) (f: option A -> option B -> option C), + f None None = None -> + forall (m1: t A) (m2: t B) (i: elt), + get i (combine f m1 m2) = f (get i m1) (get i m2). + + (** Enumerating the bindings of a tree. *) + Parameter elements: + forall (A: Type), t A -> list (elt * A). + Axiom elements_correct: + forall (A: Type) (m: t A) (i: elt) (v: A), + get i m = Some v -> In (i, v) (elements m). + Axiom elements_complete: + forall (A: Type) (m: t A) (i: elt) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Axiom elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + Axiom elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + Axiom elements_remove: + forall (A: Type) i v (m: t A), + get i m = Some v -> + exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2. + + (** Folding a function over all bindings of a tree. *) + Parameter fold: + forall (A B: Type), (B -> elt -> A -> B) -> t A -> B -> B. + Axiom fold_spec: + forall (A B: Type) (f: B -> elt -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + (** Same as [fold], but the function does not receive the [elt] argument. *) + Parameter fold1: + forall (A B: Type), (B -> A -> B) -> t A -> B -> B. + Axiom fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. +End TREE. + +(** * The abstract signatures of maps *) + +Module Type MAP. + Parameter elt: Type. + Parameter elt_eq: forall (a b: elt), {a = b} + {a <> b}. + Parameter t: Type -> Type. + Parameter init: forall (A: Type), A -> t A. + Parameter get: forall (A: Type), elt -> t A -> A. + Parameter set: forall (A: Type), elt -> A -> t A -> t A. + Axiom gi: + forall (A: Type) (i: elt) (x: A), get i (init x) = x. + Axiom gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = x. + Axiom gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Axiom gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then x else get i m. + Axiom gsident: + forall (A: Type) (i j: elt) (m: t A), get j (set i (get i m) m) = get j m. + Parameter map: forall (A B: Type), (A -> B) -> t A -> t B. + Axiom gmap: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map f m) = f(get i m). +End MAP. + +(** * An implementation of trees over type [positive] *) + +Module PTree <: TREE. + Definition elt := positive. + Definition elt_eq := peq. + + Inductive tree (A : Type) : Type := + | Leaf : tree A + | Node : tree A -> option A -> tree A -> tree A. + + Arguments Leaf {A}. + Arguments Node [A]. + Scheme tree_ind := Induction for tree Sort Prop. + + Definition t := tree. + + Definition empty (A : Type) := (Leaf : t A). + + Fixpoint get (A : Type) (i : positive) (m : t A) {struct i} : option A := + match m with + | Leaf => None + | Node l o r => + match i with + | xH => o + | xO ii => get ii l + | xI ii => get ii r + end + end. + + Fixpoint set (A : Type) (i : positive) (v : A) (m : t A) {struct i} : t A := + match m with + | Leaf => + match i with + | xH => Node Leaf (Some v) Leaf + | xO ii => Node (set ii v Leaf) None Leaf + | xI ii => Node Leaf None (set ii v Leaf) + end + | Node l o r => + match i with + | xH => Node l (Some v) r + | xO ii => Node (set ii v l) o r + | xI ii => Node l o (set ii v r) + end + end. + + Fixpoint remove (A : Type) (i : positive) (m : t A) {struct i} : t A := + match i with + | xH => + match m with + | Leaf => Leaf + | Node Leaf o Leaf => Leaf + | Node l o r => Node l None r + end + | xO ii => + match m with + | Leaf => Leaf + | Node l None Leaf => + match remove ii l with + | Leaf => Leaf + | mm => Node mm None Leaf + end + | Node l o r => Node (remove ii l) o r + end + | xI ii => + match m with + | Leaf => Leaf + | Node Leaf None r => + match remove ii r with + | Leaf => Leaf + | mm => Node Leaf None mm + end + | Node l o r => Node l o (remove ii r) + end + end. + + Theorem gempty: + forall (A: Type) (i: positive), get i (empty A) = None. + Proof. + induction i; simpl; auto. + Qed. + + Theorem gss: + forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = Some x. + Proof. + induction i; destruct m; simpl; auto. + Qed. + + Lemma gleaf : forall (A : Type) (i : positive), get i (Leaf : t A) = None. + Proof. exact gempty. Qed. + + Theorem gso: + forall (A: Type) (i j: positive) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + induction i; intros; destruct j; destruct m; simpl; + try rewrite <- (gleaf A i); auto; try apply IHi; congruence. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: positive) (x: A) (m: t A), + get i (set j x m) = if peq i j then Some x else get i m. + Proof. + intros. + destruct (peq i j); [ rewrite e; apply gss | apply gso; auto ]. + Qed. + + Theorem gsident: + forall (A: Type) (i: positive) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + Proof. + induction i; intros; destruct m; simpl; simpl in H; try congruence. + rewrite (IHi m2 v H); congruence. + rewrite (IHi m1 v H); congruence. + Qed. + + Theorem set2: + forall (A: Type) (i: elt) (m: t A) (v1 v2: A), + set i v2 (set i v1 m) = set i v2 m. + Proof. + induction i; intros; destruct m; simpl; try (rewrite IHi); auto. + Qed. + + Lemma rleaf : forall (A : Type) (i : positive), remove i (Leaf : t A) = Leaf. + Proof. destruct i; simpl; auto. Qed. + + Theorem grs: + forall (A: Type) (i: positive) (m: t A), get i (remove i m) = None. + Proof. + induction i; destruct m. + simpl; auto. + destruct m1; destruct o; destruct m2 as [ | ll oo rr]; simpl; auto. + rewrite (rleaf A i); auto. + cut (get i (remove i (Node ll oo rr)) = None). + destruct (remove i (Node ll oo rr)); auto; apply IHi. + apply IHi. + simpl; auto. + destruct m1 as [ | ll oo rr]; destruct o; destruct m2; simpl; auto. + rewrite (rleaf A i); auto. + cut (get i (remove i (Node ll oo rr)) = None). + destruct (remove i (Node ll oo rr)); auto; apply IHi. + apply IHi. + simpl; auto. + destruct m1; destruct m2; simpl; auto. + Qed. + + Theorem gro: + forall (A: Type) (i j: positive) (m: t A), + i <> j -> get i (remove j m) = get i m. + Proof. + induction i; intros; destruct j; destruct m; + try rewrite (rleaf A (xI j)); + try rewrite (rleaf A (xO j)); + try rewrite (rleaf A 1); auto; + destruct m1; destruct o; destruct m2; + simpl; + try apply IHi; try congruence; + try rewrite (rleaf A j); auto; + try rewrite (gleaf A i); auto. + cut (get i (remove j (Node m2_1 o m2_2)) = get i (Node m2_1 o m2_2)); + [ destruct (remove j (Node m2_1 o m2_2)); try rewrite (gleaf A i); auto + | apply IHi; congruence ]. + destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf A i); + auto. + destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf A i); + auto. + cut (get i (remove j (Node m1_1 o0 m1_2)) = get i (Node m1_1 o0 m1_2)); + [ destruct (remove j (Node m1_1 o0 m1_2)); try rewrite (gleaf A i); auto + | apply IHi; congruence ]. + destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf A i); + auto. + destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf A i); + auto. + Qed. + + Theorem grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + Proof. + intros. destruct (elt_eq i j). subst j. apply grs. apply gro; auto. + Qed. + + Section BOOLEAN_EQUALITY. + + Variable A: Type. + Variable beqA: A -> A -> bool. + + Fixpoint bempty (m: t A) : bool := + match m with + | Leaf => true + | Node l None r => bempty l && bempty r + | Node l (Some _) r => false + end. + + Fixpoint beq (m1 m2: t A) {struct m1} : bool := + match m1, m2 with + | Leaf, _ => bempty m2 + | _, Leaf => bempty m1 + | Node l1 o1 r1, Node l2 o2 r2 => + match o1, o2 with + | None, None => true + | Some y1, Some y2 => beqA y1 y2 + | _, _ => false + end + && beq l1 l2 && beq r1 r2 + end. + + Lemma bempty_correct: + forall m, bempty m = true <-> (forall x, get x m = None). + Proof. + induction m; simpl. + split; intros. apply gleaf. auto. + destruct o; split; intros. + congruence. + generalize (H xH); simpl; congruence. + destruct (andb_prop _ _ H). rewrite IHm1 in H0. rewrite IHm2 in H1. + destruct x; simpl; auto. + apply andb_true_intro; split. + apply IHm1. intros; apply (H (xO x)). + apply IHm2. intros; apply (H (xI x)). + Qed. + + Lemma beq_correct: + forall m1 m2, + beq m1 m2 = true <-> + (forall (x: elt), + match get x m1, get x m2 with + | None, None => True + | Some y1, Some y2 => beqA y1 y2 = true + | _, _ => False + end). + Proof. + induction m1; intros. + - simpl. rewrite bempty_correct. split; intros. + rewrite gleaf. rewrite H. auto. + generalize (H x). rewrite gleaf. destruct (get x m2); tauto. + - destruct m2. + + unfold beq. rewrite bempty_correct. split; intros. + rewrite H. rewrite gleaf. auto. + generalize (H x). rewrite gleaf. destruct (get x (Node m1_1 o m1_2)); tauto. + + simpl. split; intros. + * destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0). + rewrite IHm1_1 in H3. rewrite IHm1_2 in H1. + destruct x; simpl. apply H1. apply H3. + destruct o; destruct o0; auto || congruence. + * apply andb_true_intro. split. apply andb_true_intro. split. + generalize (H xH); simpl. destruct o; destruct o0; tauto. + apply IHm1_1. intros; apply (H (xO x)). + apply IHm1_2. intros; apply (H (xI x)). + Qed. + + End BOOLEAN_EQUALITY. + + Fixpoint prev_append (i j: positive) {struct i} : positive := + match i with + | xH => j + | xI i' => prev_append i' (xI j) + | xO i' => prev_append i' (xO j) + end. + + Definition prev (i: positive) : positive := + prev_append i xH. + + Lemma prev_append_prev i j: + prev (prev_append i j) = prev_append j i. + Proof. + revert j. unfold prev. + induction i as [i IH|i IH|]. 3: reflexivity. + intros j. simpl. rewrite IH. reflexivity. + intros j. simpl. rewrite IH. reflexivity. + Qed. + + Lemma prev_involutive i : + prev (prev i) = i. + Proof (prev_append_prev i xH). + + Lemma prev_append_inj i j j' : + prev_append i j = prev_append i j' -> j = j'. + Proof. + revert j j'. + induction i as [i Hi|i Hi|]; intros j j' H; auto; + specialize (Hi _ _ H); congruence. + Qed. + + Fixpoint xmap (A B : Type) (f : positive -> A -> B) (m : t A) (i : positive) + {struct m} : t B := + match m with + | Leaf => Leaf + | Node l o r => Node (xmap f l (xO i)) + (match o with None => None | Some x => Some (f (prev i) x) end) + (xmap f r (xI i)) + end. + + Definition map (A B : Type) (f : positive -> A -> B) m := xmap f m xH. + + Lemma xgmap: + forall (A B: Type) (f: positive -> A -> B) (i j : positive) (m: t A), + get i (xmap f m j) = option_map (f (prev (prev_append i j))) (get i m). + Proof. + induction i; intros; destruct m; simpl; auto. + Qed. + + Theorem gmap: + forall (A B: Type) (f: positive -> A -> B) (i: positive) (m: t A), + get i (map f m) = option_map (f i) (get i m). + Proof. + intros A B f i m. + unfold map. + rewrite xgmap. repeat f_equal. exact (prev_involutive i). + Qed. + + Fixpoint map1 (A B: Type) (f: A -> B) (m: t A) {struct m} : t B := + match m with + | Leaf => Leaf + | Node l o r => Node (map1 f l) (option_map f o) (map1 f r) + end. + + Theorem gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + Proof. + induction i; intros; destruct m; simpl; auto. + Qed. + + Definition Node' (A: Type) (l: t A) (x: option A) (r: t A): t A := + match l, x, r with + | Leaf, None, Leaf => Leaf + | _, _, _ => Node l x r + end. + + Lemma gnode': + forall (A: Type) (l r: t A) (x: option A) (i: positive), + get i (Node' l x r) = get i (Node l x r). + Proof. + intros. unfold Node'. + destruct l; destruct x; destruct r; auto. + destruct i; simpl; auto; rewrite gleaf; auto. + Qed. + + Fixpoint filter1 (A: Type) (pred: A -> bool) (m: t A) {struct m} : t A := + match m with + | Leaf => Leaf + | Node l o r => + let o' := match o with None => None | Some x => if pred x then o else None end in + Node' (filter1 pred l) o' (filter1 pred r) + end. + + Theorem gfilter1: + forall (A: Type) (pred: A -> bool) (i: elt) (m: t A), + get i (filter1 pred m) = + match get i m with None => None | Some x => if pred x then Some x else None end. + Proof. + intros until m. revert m i. induction m; simpl; intros. + rewrite gleaf; auto. + rewrite gnode'. destruct i; simpl; auto. destruct o; auto. + Qed. + + Section COMBINE. + + Variables A B C: Type. + Variable f: option A -> option B -> option C. + Hypothesis f_none_none: f None None = None. + + Fixpoint xcombine_l (m : t A) {struct m} : t C := + match m with + | Leaf => Leaf + | Node l o r => Node' (xcombine_l l) (f o None) (xcombine_l r) + end. + + Lemma xgcombine_l : + forall (m: t A) (i : positive), + get i (xcombine_l m) = f (get i m) None. + Proof. + induction m; intros; simpl. + repeat rewrite gleaf. auto. + rewrite gnode'. destruct i; simpl; auto. + Qed. + + Fixpoint xcombine_r (m : t B) {struct m} : t C := + match m with + | Leaf => Leaf + | Node l o r => Node' (xcombine_r l) (f None o) (xcombine_r r) + end. + + Lemma xgcombine_r : + forall (m: t B) (i : positive), + get i (xcombine_r m) = f None (get i m). + Proof. + induction m; intros; simpl. + repeat rewrite gleaf. auto. + rewrite gnode'. destruct i; simpl; auto. + Qed. + + Fixpoint combine (m1: t A) (m2: t B) {struct m1} : t C := + match m1 with + | Leaf => xcombine_r m2 + | Node l1 o1 r1 => + match m2 with + | Leaf => xcombine_l m1 + | Node l2 o2 r2 => Node' (combine l1 l2) (f o1 o2) (combine r1 r2) + end + end. + + Theorem gcombine: + forall (m1: t A) (m2: t B) (i: positive), + get i (combine m1 m2) = f (get i m1) (get i m2). + Proof. + induction m1; intros; simpl. + rewrite gleaf. apply xgcombine_r. + destruct m2; simpl. + rewrite gleaf. rewrite <- xgcombine_l. auto. + repeat rewrite gnode'. destruct i; simpl; auto. + Qed. + + End COMBINE. + + Lemma xcombine_lr : + forall (A B: Type) (f g : option A -> option A -> option B) (m : t A), + (forall (i j : option A), f i j = g j i) -> + xcombine_l f m = xcombine_r g m. + Proof. + induction m; intros; simpl; auto. + rewrite IHm1; auto. + rewrite IHm2; auto. + rewrite H; auto. + Qed. + + Theorem combine_commut: + forall (A B: Type) (f g: option A -> option A -> option B), + (forall (i j: option A), f i j = g j i) -> + forall (m1 m2: t A), + combine f m1 m2 = combine g m2 m1. + Proof. + intros A B f g EQ1. + assert (EQ2: forall (i j: option A), g i j = f j i). + intros; auto. + induction m1; intros; destruct m2; simpl; + try rewrite EQ1; + repeat rewrite (xcombine_lr f g); + repeat rewrite (xcombine_lr g f); + auto. + rewrite IHm1_1. + rewrite IHm1_2. + auto. + Qed. + + Fixpoint xelements (A : Type) (m : t A) (i : positive) + (k: list (positive * A)) {struct m} + : list (positive * A) := + match m with + | Leaf => k + | Node l None r => + xelements l (xO i) (xelements r (xI i) k) + | Node l (Some x) r => + xelements l (xO i) + ((prev i, x) :: xelements r (xI i) k) + end. + + Definition elements (A: Type) (m : t A) := xelements m xH nil. + + Remark xelements_append: + forall A (m: t A) i k1 k2, + xelements m i (k1 ++ k2) = xelements m i k1 ++ k2. + Proof. + induction m; intros; simpl. + - auto. + - destruct o; rewrite IHm2; rewrite <- IHm1; auto. + Qed. + + Remark xelements_leaf: + forall A i, xelements (@Leaf A) i nil = nil. + Proof. + intros; reflexivity. + Qed. + + Remark xelements_node: + forall A (m1: t A) o (m2: t A) i, + xelements (Node m1 o m2) i nil = + xelements m1 (xO i) nil + ++ match o with None => nil | Some v => (prev i, v) :: nil end + ++ xelements m2 (xI i) nil. + Proof. + intros. simpl. destruct o; simpl; rewrite <- xelements_append; auto. + Qed. + + Lemma xelements_incl: + forall (A: Type) (m: t A) (i : positive) k x, + In x k -> In x (xelements m i k). + Proof. + induction m; intros; simpl. + auto. + destruct o. + apply IHm1. simpl; right; auto. + auto. + Qed. + + Lemma xelements_correct: + forall (A: Type) (m: t A) (i j : positive) (v: A) k, + get i m = Some v -> In (prev (prev_append i j), v) (xelements m j k). + Proof. + induction m; intros. + rewrite (gleaf A i) in H; congruence. + destruct o; destruct i; simpl; simpl in H. + apply xelements_incl. right. auto. + auto. + inv H. apply xelements_incl. left. reflexivity. + apply xelements_incl. auto. + auto. + inv H. + Qed. + + Theorem elements_correct: + forall (A: Type) (m: t A) (i: positive) (v: A), + get i m = Some v -> In (i, v) (elements m). + Proof. + intros A m i v H. + generalize (xelements_correct m i xH nil H). rewrite prev_append_prev. exact id. + Qed. + + Lemma in_xelements: + forall (A: Type) (m: t A) (i k: positive) (v: A) , + In (k, v) (xelements m i nil) -> + exists j, k = prev (prev_append j i) /\ get j m = Some v. + Proof. + induction m; intros. + - rewrite xelements_leaf in H. contradiction. + - rewrite xelements_node in H. rewrite ! in_app_iff in H. destruct H as [P | [P | P]]. + + exploit IHm1; eauto. intros (j & Q & R). exists (xO j); auto. + + destruct o; simpl in P; intuition auto. inv H. exists xH; auto. + + exploit IHm2; eauto. intros (j & Q & R). exists (xI j); auto. + Qed. + + Theorem elements_complete: + forall (A: Type) (m: t A) (i: positive) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Proof. + unfold elements. intros A m i v H. exploit in_xelements; eauto. intros (j & P & Q). + rewrite prev_append_prev in P. change i with (prev_append 1 i) in P. + exploit prev_append_inj; eauto. intros; congruence. + Qed. + + Definition xkeys (A: Type) (m: t A) (i: positive) := + List.map (@fst positive A) (xelements m i nil). + + Remark xkeys_leaf: + forall A i, xkeys (@Leaf A) i = nil. + Proof. + intros; reflexivity. + Qed. + + Remark xkeys_node: + forall A (m1: t A) o (m2: t A) i, + xkeys (Node m1 o m2) i = + xkeys m1 (xO i) + ++ match o with None => nil | Some v => prev i :: nil end + ++ xkeys m2 (xI i). + Proof. + intros. unfold xkeys. rewrite xelements_node. rewrite ! map_app. destruct o; auto. + Qed. + + Lemma in_xkeys: + forall (A: Type) (m: t A) (i k: positive), + In k (xkeys m i) -> + (exists j, k = prev (prev_append j i)). + Proof. + unfold xkeys; intros. + apply (list_in_map_inv) in H. destruct H as ((j, v) & -> & H). + exploit in_xelements; eauto. intros (k & P & Q). exists k; auto. + Qed. + + Lemma xelements_keys_norepet: + forall (A: Type) (m: t A) (i: positive), + list_norepet (xkeys m i). + Proof. + induction m; intros. + - rewrite xkeys_leaf; constructor. + - assert (NOTIN1: ~ In (prev i) (xkeys m1 (xO i))). + { red; intros. exploit in_xkeys; eauto. intros (j & EQ). + rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. } + assert (NOTIN2: ~ In (prev i) (xkeys m2 (xI i))). + { red; intros. exploit in_xkeys; eauto. intros (j & EQ). + rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. } + assert (DISJ: forall x, In x (xkeys m1 (xO i)) -> In x (xkeys m2 (xI i)) -> False). + { intros. exploit in_xkeys. eexact H. intros (j1 & EQ1). + exploit in_xkeys. eexact H0. intros (j2 & EQ2). + rewrite prev_append_prev in *. simpl in *. rewrite EQ2 in EQ1. apply prev_append_inj in EQ1. discriminate. } + rewrite xkeys_node. apply list_norepet_append. auto. + destruct o; simpl; auto. constructor; auto. + red; intros. red; intros; subst y. destruct o; simpl in H0. + destruct H0. subst x. tauto. eauto. eauto. + Qed. + + Theorem elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + Proof. + intros. apply (xelements_keys_norepet m xH). + Qed. + + Remark xelements_empty: + forall (A: Type) (m: t A) i, (forall i, get i m = None) -> xelements m i nil = nil. + Proof. + induction m; intros. + auto. + rewrite xelements_node. rewrite IHm1, IHm2. destruct o; auto. + generalize (H xH); simpl; congruence. + intros. apply (H (xI i0)). + intros. apply (H (xO i0)). + Qed. + + Theorem elements_canonical_order': + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i, option_rel R (get i m) (get i n)) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + Proof. + intros until n. unfold elements. generalize 1%positive. revert m n. + induction m; intros. + - simpl. rewrite xelements_empty. constructor. + intros. specialize (H i). rewrite gempty in H. inv H; auto. + - destruct n as [ | n1 o' n2 ]. + + rewrite (xelements_empty (Node m1 o m2)). simpl; constructor. + intros. specialize (H i). rewrite gempty in H. inv H; auto. + + rewrite ! xelements_node. repeat apply list_forall2_app. + apply IHm1. intros. apply (H (xO i)). + generalize (H xH); simpl; intros OR; inv OR. + constructor. + constructor. auto. constructor. + apply IHm2. intros. apply (H (xI i)). + Qed. + + Theorem elements_canonical_order: + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i x, get i m = Some x -> exists y, get i n = Some y /\ R x y) -> + (forall i y, get i n = Some y -> exists x, get i m = Some x /\ R x y) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + Proof. + intros. apply elements_canonical_order'. + intros. destruct (get i m) as [x|] eqn:GM. + exploit H; eauto. intros (y & P & Q). rewrite P; constructor; auto. + destruct (get i n) as [y|] eqn:GN. + exploit H0; eauto. intros (x & P & Q). congruence. + constructor. + Qed. + + Theorem elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + Proof. + intros. + exploit (@elements_canonical_order' _ _ (fun (x y: A) => x = y) m n). + intros. rewrite H. destruct (get i n); constructor; auto. + induction 1. auto. destruct a1 as [a2 a3]; destruct b1 as [b2 b3]; simpl in *. + destruct H0. congruence. + Qed. + + Lemma xelements_remove: + forall (A: Type) v (m: t A) i j, + get i m = Some v -> + exists l1 l2, + xelements m j nil = l1 ++ (prev (prev_append i j), v) :: l2 + /\ xelements (remove i m) j nil = l1 ++ l2. + Proof. + induction m; intros. + - rewrite gleaf in H; discriminate. + - assert (REMOVE: xelements (remove i (Node m1 o m2)) j nil = + xelements (match i with + | xH => Node m1 None m2 + | xO ii => Node (remove ii m1) o m2 + | xI ii => Node m1 o (remove ii m2) end) + j nil). + { + destruct i; simpl remove. + destruct m1; auto. destruct o; auto. destruct (remove i m2); auto. + destruct o; auto. destruct m2; auto. destruct (remove i m1); auto. + destruct m1; auto. destruct m2; auto. + } + rewrite REMOVE. destruct i; simpl in H. + + destruct (IHm2 i (xI j) H) as (l1 & l2 & EQ & EQ'). + exists (xelements m1 (xO j) nil ++ + match o with None => nil | Some x => (prev j, x) :: nil end ++ + l1); + exists l2; split. + rewrite xelements_node, EQ, ! app_ass. auto. + rewrite xelements_node, EQ', ! app_ass. auto. + + destruct (IHm1 i (xO j) H) as (l1 & l2 & EQ & EQ'). + exists l1; + exists (l2 ++ + match o with None => nil | Some x => (prev j, x) :: nil end ++ + xelements m2 (xI j) nil); + split. + rewrite xelements_node, EQ, ! app_ass. auto. + rewrite xelements_node, EQ', ! app_ass. auto. + + subst o. exists (xelements m1 (xO j) nil); exists (xelements m2 (xI j) nil); split. + rewrite xelements_node. rewrite prev_append_prev. auto. + rewrite xelements_node; auto. + Qed. + + Theorem elements_remove: + forall (A: Type) i v (m: t A), + get i m = Some v -> + exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2. + Proof. + intros. exploit xelements_remove. eauto. instantiate (1 := xH). + rewrite prev_append_prev. auto. + Qed. + + Fixpoint xfold (A B: Type) (f: B -> positive -> A -> B) + (i: positive) (m: t A) (v: B) {struct m} : B := + match m with + | Leaf => v + | Node l None r => + let v1 := xfold f (xO i) l v in + xfold f (xI i) r v1 + | Node l (Some x) r => + let v1 := xfold f (xO i) l v in + let v2 := f v1 (prev i) x in + xfold f (xI i) r v2 + end. + + Definition fold (A B : Type) (f: B -> positive -> A -> B) (m: t A) (v: B) := + xfold f xH m v. + + Lemma xfold_xelements: + forall (A B: Type) (f: B -> positive -> A -> B) m i v l, + List.fold_left (fun a p => f a (fst p) (snd p)) l (xfold f i m v) = + List.fold_left (fun a p => f a (fst p) (snd p)) (xelements m i l) v. + Proof. + induction m; intros. + simpl. auto. + destruct o; simpl. + rewrite <- IHm1. simpl. rewrite <- IHm2. auto. + rewrite <- IHm1. rewrite <- IHm2. auto. + Qed. + + Theorem fold_spec: + forall (A B: Type) (f: B -> positive -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + Proof. + intros. unfold fold, elements. rewrite <- xfold_xelements. auto. + Qed. + + Fixpoint fold1 (A B: Type) (f: B -> A -> B) (m: t A) (v: B) {struct m} : B := + match m with + | Leaf => v + | Node l None r => + let v1 := fold1 f l v in + fold1 f r v1 + | Node l (Some x) r => + let v1 := fold1 f l v in + let v2 := f v1 x in + fold1 f r v2 + end. + + Lemma fold1_xelements: + forall (A B: Type) (f: B -> A -> B) m i v l, + List.fold_left (fun a p => f a (snd p)) l (fold1 f m v) = + List.fold_left (fun a p => f a (snd p)) (xelements m i l) v. + Proof. + induction m; intros. + simpl. auto. + destruct o; simpl. + rewrite <- IHm1. simpl. rewrite <- IHm2. auto. + rewrite <- IHm1. rewrite <- IHm2. auto. + Qed. + + Theorem fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. + Proof. + intros. apply fold1_xelements with (l := @nil (positive * A)). + Qed. + +End PTree. + +(** * An implementation of maps over type [positive] *) + +Module PMap <: MAP. + Definition elt := positive. + Definition elt_eq := peq. + + Definition t (A : Type) : Type := (A * PTree.t A)%type. + + Definition init (A : Type) (x : A) := + (x, PTree.empty A). + + Definition get (A : Type) (i : positive) (m : t A) := + match PTree.get i (snd m) with + | Some x => x + | None => fst m + end. + + Definition set (A : Type) (i : positive) (x : A) (m : t A) := + (fst m, PTree.set i x (snd m)). + + Theorem gi: + forall (A: Type) (i: positive) (x: A), get i (init x) = x. + Proof. + intros. unfold init. unfold get. simpl. rewrite PTree.gempty. auto. + Qed. + + Theorem gss: + forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = x. + Proof. + intros. unfold get. unfold set. simpl. rewrite PTree.gss. auto. + Qed. + + Theorem gso: + forall (A: Type) (i j: positive) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. unfold get. unfold set. simpl. rewrite PTree.gso; auto. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: positive) (x: A) (m: t A), + get i (set j x m) = if peq i j then x else get i m. + Proof. + intros. destruct (peq i j). + rewrite e. apply gss. auto. + apply gso. auto. + Qed. + + Theorem gsident: + forall (A: Type) (i j: positive) (m: t A), + get j (set i (get i m) m) = get j m. + Proof. + intros. destruct (peq i j). + rewrite e. rewrite gss. auto. + rewrite gso; auto. + Qed. + + Definition map (A B : Type) (f : A -> B) (m : t A) : t B := + (f (fst m), PTree.map1 f (snd m)). + + Theorem gmap: + forall (A B: Type) (f: A -> B) (i: positive) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold map. unfold get. simpl. rewrite PTree.gmap1. + unfold option_map. destruct (PTree.get i (snd m)); auto. + Qed. + + Theorem set2: + forall (A: Type) (i: elt) (x y: A) (m: t A), + set i y (set i x m) = set i y m. + Proof. + intros. unfold set. simpl. decEq. apply PTree.set2. + Qed. + +End PMap. + +(** * An implementation of maps over any type that injects into type [positive] *) + +Module Type INDEXED_TYPE. + Parameter t: Type. + Parameter index: t -> positive. + Axiom index_inj: forall (x y: t), index x = index y -> x = y. + Parameter eq: forall (x y: t), {x = y} + {x <> y}. +End INDEXED_TYPE. + +Module IMap(X: INDEXED_TYPE). + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t : Type -> Type := PMap.t. + Definition init (A: Type) (x: A) := PMap.init x. + Definition get (A: Type) (i: X.t) (m: t A) := PMap.get (X.index i) m. + Definition set (A: Type) (i: X.t) (v: A) (m: t A) := PMap.set (X.index i) v m. + Definition map (A B: Type) (f: A -> B) (m: t A) : t B := PMap.map f m. + + Lemma gi: + forall (A: Type) (x: A) (i: X.t), get i (init x) = x. + Proof. + intros. unfold get, init. apply PMap.gi. + Qed. + + Lemma gss: + forall (A: Type) (i: X.t) (x: A) (m: t A), get i (set i x m) = x. + Proof. + intros. unfold get, set. apply PMap.gss. + Qed. + + Lemma gso: + forall (A: Type) (i j: X.t) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. unfold get, set. apply PMap.gso. + red. intro. apply H. apply X.index_inj; auto. + Qed. + + Lemma gsspec: + forall (A: Type) (i j: X.t) (x: A) (m: t A), + get i (set j x m) = if X.eq i j then x else get i m. + Proof. + intros. unfold get, set. + rewrite PMap.gsspec. + case (X.eq i j); intro. + subst j. rewrite peq_true. reflexivity. + rewrite peq_false. reflexivity. + red; intro. elim n. apply X.index_inj; auto. + Qed. + + Lemma gmap: + forall (A B: Type) (f: A -> B) (i: X.t) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold map, get. apply PMap.gmap. + Qed. + + Lemma set2: + forall (A: Type) (i: elt) (x y: A) (m: t A), + set i y (set i x m) = set i y m. + Proof. + intros. unfold set. apply PMap.set2. + Qed. + +End IMap. + +Module ZIndexed. + Definition t := Z. + Definition index (z: Z): positive := + match z with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + Lemma index_inj: forall (x y: Z), index x = index y -> x = y. + Proof. + unfold index; destruct x; destruct y; intros; + try discriminate; try reflexivity. + congruence. + congruence. + Qed. + Definition eq := zeq. +End ZIndexed. + +Module ZMap := IMap(ZIndexed). + +Module NIndexed. + Definition t := N. + Definition index (n: N): positive := + match n with + | N0 => xH + | Npos p => xO p + end. + Lemma index_inj: forall (x y: N), index x = index y -> x = y. + Proof. + unfold index; destruct x; destruct y; intros; + try discriminate; try reflexivity. + congruence. + Qed. + Lemma eq: forall (x y: N), {x = y} + {x <> y}. + Proof. + decide equality. apply peq. + Qed. +End NIndexed. + +Module NMap := IMap(NIndexed). + +(** * An implementation of maps over any type with decidable equality *) + +Module Type EQUALITY_TYPE. + Parameter t: Type. + Parameter eq: forall (x y: t), {x = y} + {x <> y}. +End EQUALITY_TYPE. + +Module EMap(X: EQUALITY_TYPE) <: MAP. + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t (A: Type) := X.t -> A. + Definition init (A: Type) (v: A) := fun (_: X.t) => v. + Definition get (A: Type) (x: X.t) (m: t A) := m x. + Definition set (A: Type) (x: X.t) (v: A) (m: t A) := + fun (y: X.t) => if X.eq y x then v else m y. + Lemma gi: + forall (A: Type) (i: elt) (x: A), init x i = x. + Proof. + intros. reflexivity. + Qed. + Lemma gss: + forall (A: Type) (i: elt) (x: A) (m: t A), (set i x m) i = x. + Proof. + intros. unfold set. case (X.eq i i); intro. + reflexivity. tauto. + Qed. + Lemma gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> (set j x m) i = m i. + Proof. + intros. unfold set. case (X.eq i j); intro. + congruence. reflexivity. + Qed. + Lemma gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then x else get i m. + Proof. + intros. unfold get, set, elt_eq. reflexivity. + Qed. + Lemma gsident: + forall (A: Type) (i j: elt) (m: t A), get j (set i (get i m) m) = get j m. + Proof. + intros. unfold get, set. case (X.eq j i); intro. + congruence. reflexivity. + Qed. + Definition map (A B: Type) (f: A -> B) (m: t A) := + fun (x: X.t) => f(m x). + Lemma gmap: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold get, map. reflexivity. + Qed. +End EMap. + +(** * A partial implementation of trees over any type that injects into type [positive] *) + +Module ITree(X: INDEXED_TYPE). + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t : Type -> Type := PTree.t. + + Definition empty (A: Type): t A := PTree.empty A. + Definition get (A: Type) (k: elt) (m: t A): option A := PTree.get (X.index k) m. + Definition set (A: Type) (k: elt) (v: A) (m: t A): t A := PTree.set (X.index k) v m. + Definition remove (A: Type) (k: elt) (m: t A): t A := PTree.remove (X.index k) m. + + Theorem gempty: + forall (A: Type) (i: elt), get i (empty A) = None. + Proof. + intros. apply PTree.gempty. + Qed. + Theorem gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = Some x. + Proof. + intros. apply PTree.gss. + Qed. + Theorem gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. apply PTree.gso. red; intros; elim H; apply X.index_inj; auto. + Qed. + Theorem gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then Some x else get i m. + Proof. + intros. destruct (elt_eq i j). subst j; apply gss. apply gso; auto. + Qed. + Theorem grs: + forall (A: Type) (i: elt) (m: t A), get i (remove i m) = None. + Proof. + intros. apply PTree.grs. + Qed. + Theorem gro: + forall (A: Type) (i j: elt) (m: t A), + i <> j -> get i (remove j m) = get i m. + Proof. + intros. apply PTree.gro. red; intros; elim H; apply X.index_inj; auto. + Qed. + Theorem grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + Proof. + intros. destruct (elt_eq i j). subst j; apply grs. apply gro; auto. + Qed. + + Definition beq: forall (A: Type), (A -> A -> bool) -> t A -> t A -> bool := PTree.beq. + Theorem beq_sound: + forall (A: Type) (eqA: A -> A -> bool) (t1 t2: t A), + beq eqA t1 t2 = true -> + forall (x: elt), + match get x t1, get x t2 with + | None, None => True + | Some y1, Some y2 => eqA y1 y2 = true + | _, _ => False + end. + Proof. + unfold beq, get. intros. rewrite PTree.beq_correct in H. apply H. + Qed. + + Definition combine: forall (A B C: Type), (option A -> option B -> option C) -> t A -> t B -> t C := PTree.combine. + Theorem gcombine: + forall (A B C: Type) (f: option A -> option B -> option C), + f None None = None -> + forall (m1: t A) (m2: t B) (i: elt), + get i (combine f m1 m2) = f (get i m1) (get i m2). + Proof. + intros. apply PTree.gcombine. auto. + Qed. +End ITree. + +Module ZTree := ITree(ZIndexed). + +(** * Additional properties over trees *) + +Module Tree_Properties(T: TREE). + +(** Two induction principles over [fold]. *) + +Section TREE_FOLD_IND. + +Variables V A: Type. +Variable f: A -> T.elt -> V -> A. +Variable P: T.t V -> A -> Type. +Variable init: A. +Variable m_final: T.t V. + +Hypothesis H_base: + forall m, + (forall k, T.get k m = None) -> + P m init. + +Hypothesis H_rec: + forall m a k v, + T.get k m = Some v -> T.get k m_final = Some v -> + P (T.remove k m) a -> P m (f a k v). + +Let f' (p : T.elt * V) (a: A) := f a (fst p) (snd p). + +Let P' (l: list (T.elt * V)) (a: A) : Type := + forall m, (forall k v, In (k, v) l <-> T.get k m = Some v) -> P m a. + +Let H_base': + P' nil init. +Proof. + intros m EQV. apply H_base. + intros. destruct (T.get k m) as [v|] eqn:G; auto. + apply EQV in G. contradiction. +Qed. + +Let H_rec': + forall k v l a, + ~In k (List.map fst l) -> + T.get k m_final = Some v -> + P' l a -> + P' ((k, v) :: l) (f a k v). +Proof. + unfold P'; intros k v l a NOTIN FINAL HR m EQV. + set (m0 := T.remove k m). + apply H_rec. +- apply EQV. simpl; auto. +- auto. +- apply HR. intros k' v'. rewrite T.grspec. split; intros; destruct (T.elt_eq k' k). + + subst k'. elim NOTIN. change k with (fst (k, v')). apply List.in_map; auto. + + apply EQV. simpl; auto. + + congruence. + + apply EQV in H. simpl in H. intuition congruence. +Qed. + +Lemma fold_ind_aux: + forall l, + (forall k v, In (k, v) l -> T.get k m_final = Some v) -> + list_norepet (List.map fst l) -> + P' l (List.fold_right f' init l). +Proof. + induction l as [ | [k v] l ]; simpl; intros FINAL NOREPET. +- apply H_base'. +- apply H_rec'. + + inv NOREPET. auto. + + apply FINAL. auto. + + apply IHl. auto. inv NOREPET; auto. +Defined. + +Theorem fold_ind: + P m_final (T.fold f m_final init). +Proof. + intros. + set (l' := List.rev (T.elements m_final)). + assert (P' l' (List.fold_right f' init l')). + { apply fold_ind_aux. + intros. apply T.elements_complete. apply List.in_rev. auto. + unfold l'; rewrite List.map_rev. apply list_norepet_rev. apply T.elements_keys_norepet. } + unfold l', f' in X; rewrite fold_left_rev_right in X. + rewrite T.fold_spec. apply X. + intros; simpl. rewrite <- List.in_rev. + split. apply T.elements_complete. apply T.elements_correct. +Defined. + +End TREE_FOLD_IND. + +Section TREE_FOLD_REC. + +Variables V A: Type. +Variable f: A -> T.elt -> V -> A. +Variable P: T.t V -> A -> Prop. +Variable init: A. +Variable m_final: T.t V. + +Hypothesis P_compat: + forall m m' a, + (forall x, T.get x m = T.get x m') -> + P m a -> P m' a. + +Hypothesis H_base: + P (T.empty _) init. + +Hypothesis H_rec: + forall m a k v, + T.get k m = None -> T.get k m_final = Some v -> P m a -> P (T.set k v m) (f a k v). + +Theorem fold_rec: + P m_final (T.fold f m_final init). +Proof. + apply fold_ind. +- intros. apply P_compat with (T.empty V); auto. + + intros. rewrite T.gempty. auto. +- intros. apply P_compat with (T.set k v (T.remove k m)). + + intros. rewrite T.gsspec, T.grspec. destruct (T.elt_eq x k); auto. congruence. + + apply H_rec; auto. apply T.grs. +Qed. + +End TREE_FOLD_REC. + +(** A nonnegative measure over trees *) + +Section MEASURE. + +Variable V: Type. + +Definition cardinal (x: T.t V) : nat := List.length (T.elements x). + +Theorem cardinal_remove: + forall x m y, T.get x m = Some y -> (cardinal (T.remove x m) < cardinal m)%nat. +Proof. + unfold cardinal; intros. + exploit T.elements_remove; eauto. intros (l1 & l2 & P & Q). + rewrite P, Q. rewrite ! app_length. simpl. lia. +Qed. + +Theorem cardinal_set: + forall x m y, T.get x m = None -> (cardinal m < cardinal (T.set x y m))%nat. +Proof. + intros. set (m' := T.set x y m). + replace (cardinal m) with (cardinal (T.remove x m')). + apply cardinal_remove with y. unfold m'; apply T.gss. + unfold cardinal. f_equal. apply T.elements_extensional. + intros. unfold m'. rewrite T.grspec, T.gsspec. + destruct (T.elt_eq i x); auto. congruence. +Qed. + +End MEASURE. + +(** Forall and exists *) + +Section FORALL_EXISTS. + +Variable A: Type. + +Definition for_all (m: T.t A) (f: T.elt -> A -> bool) : bool := + T.fold (fun b x a => b && f x a) m true. + +Lemma for_all_correct: + forall m f, + for_all m f = true <-> (forall x a, T.get x m = Some a -> f x a = true). +Proof. + intros m0 f. + unfold for_all. apply fold_rec; intros. +- (* Extensionality *) + rewrite H0. split; intros. rewrite <- H in H2; auto. rewrite H in H2; auto. +- (* Base case *) + split; intros. rewrite T.gempty in H0; congruence. auto. +- (* Inductive case *) + split; intros. + destruct (andb_prop _ _ H2). rewrite T.gsspec in H3. destruct (T.elt_eq x k). + inv H3. auto. + apply H1; auto. + apply andb_true_intro. split. + rewrite H1. intros. apply H2. rewrite T.gso; auto. congruence. + apply H2. apply T.gss. +Qed. + +Definition exists_ (m: T.t A) (f: T.elt -> A -> bool) : bool := + T.fold (fun b x a => b || f x a) m false. + +Lemma exists_correct: + forall m f, + exists_ m f = true <-> (exists x a, T.get x m = Some a /\ f x a = true). +Proof. + intros m0 f. + unfold exists_. apply fold_rec; intros. +- (* Extensionality *) + rewrite H0. split; intros (x0 & a0 & P & Q); exists x0; exists a0; split; auto; congruence. +- (* Base case *) + split; intros. congruence. destruct H as (x & a & P & Q). rewrite T.gempty in P; congruence. +- (* Inductive case *) + split; intros. + destruct (orb_true_elim _ _ H2). + rewrite H1 in e. destruct e as (x1 & a1 & P & Q). + exists x1; exists a1; split; auto. rewrite T.gso; auto. congruence. + exists k; exists v; split; auto. apply T.gss. + destruct H2 as (x1 & a1 & P & Q). apply orb_true_intro. + rewrite T.gsspec in P. destruct (T.elt_eq x1 k). + inv P. right; auto. + left. apply H1. exists x1; exists a1; auto. +Qed. + +Remark exists_for_all: + forall m f, + exists_ m f = negb (for_all m (fun x a => negb (f x a))). +Proof. + intros. unfold exists_, for_all. rewrite ! T.fold_spec. + change false with (negb true). generalize (T.elements m) true. + induction l; simpl; intros. + auto. + rewrite <- IHl. f_equal. + destruct b; destruct (f (fst a) (snd a)); reflexivity. +Qed. + +Remark for_all_exists: + forall m f, + for_all m f = negb (exists_ m (fun x a => negb (f x a))). +Proof. + intros. unfold exists_, for_all. rewrite ! T.fold_spec. + change true with (negb false). generalize (T.elements m) false. + induction l; simpl; intros. + auto. + rewrite <- IHl. f_equal. + destruct b; destruct (f (fst a) (snd a)); reflexivity. +Qed. + +Lemma for_all_false: + forall m f, + for_all m f = false <-> (exists x a, T.get x m = Some a /\ f x a = false). +Proof. + intros. rewrite for_all_exists. + rewrite negb_false_iff. rewrite exists_correct. + split; intros (x & a & P & Q); exists x; exists a; split; auto. + rewrite negb_true_iff in Q. auto. + rewrite Q; auto. +Qed. + +Lemma exists_false: + forall m f, + exists_ m f = false <-> (forall x a, T.get x m = Some a -> f x a = false). +Proof. + intros. rewrite exists_for_all. + rewrite negb_false_iff. rewrite for_all_correct. + split; intros. apply H in H0. rewrite negb_true_iff in H0. auto. rewrite H; auto. +Qed. + +End FORALL_EXISTS. + +(** More about [beq] *) + +Section BOOLEAN_EQUALITY. + +Variable A: Type. +Variable beqA: A -> A -> bool. + +Theorem beq_false: + forall m1 m2, + T.beq beqA m1 m2 = false <-> + exists x, match T.get x m1, T.get x m2 with + | None, None => False + | Some a1, Some a2 => beqA a1 a2 = false + | _, _ => True + end. +Proof. + intros; split; intros. +- (* beq = false -> existence *) + set (p1 := fun x a1 => match T.get x m2 with None => false | Some a2 => beqA a1 a2 end). + set (p2 := fun x a2 => match T.get x m1 with None => false | Some a1 => beqA a1 a2 end). + destruct (for_all m1 p1) eqn:F1; [destruct (for_all m2 p2) eqn:F2 | idtac]. + + cut (T.beq beqA m1 m2 = true). congruence. + rewrite for_all_correct in *. rewrite T.beq_correct; intros. + destruct (T.get x m1) as [a1|] eqn:X1. + generalize (F1 _ _ X1). unfold p1. destruct (T.get x m2); congruence. + destruct (T.get x m2) as [a2|] eqn:X2; auto. + generalize (F2 _ _ X2). unfold p2. rewrite X1. congruence. + + rewrite for_all_false in F2. destruct F2 as (x & a & P & Q). + exists x. rewrite P. unfold p2 in Q. destruct (T.get x m1); auto. + + rewrite for_all_false in F1. destruct F1 as (x & a & P & Q). + exists x. rewrite P. unfold p1 in Q. destruct (T.get x m2); auto. +- (* existence -> beq = false *) + destruct H as [x P]. + destruct (T.beq beqA m1 m2) eqn:E; auto. + rewrite T.beq_correct in E. + generalize (E x). destruct (T.get x m1); destruct (T.get x m2); tauto || congruence. +Qed. + +End BOOLEAN_EQUALITY. + +(** Extensional equality between trees *) + +Section EXTENSIONAL_EQUALITY. + +Variable A: Type. +Variable eqA: A -> A -> Prop. +Hypothesis eqAeq: Equivalence eqA. + +Definition Equal (m1 m2: T.t A) : Prop := + forall x, match T.get x m1, T.get x m2 with + | None, None => True + | Some a1, Some a2 => a1 === a2 + | _, _ => False + end. + +Lemma Equal_refl: forall m, Equal m m. +Proof. + intros; red; intros. destruct (T.get x m); auto. reflexivity. +Qed. + +Lemma Equal_sym: forall m1 m2, Equal m1 m2 -> Equal m2 m1. +Proof. + intros; red; intros. generalize (H x). destruct (T.get x m1); destruct (T.get x m2); auto. intros; symmetry; auto. +Qed. + +Lemma Equal_trans: forall m1 m2 m3, Equal m1 m2 -> Equal m2 m3 -> Equal m1 m3. +Proof. + intros; red; intros. generalize (H x) (H0 x). + destruct (T.get x m1); destruct (T.get x m2); try tauto; + destruct (T.get x m3); try tauto. + intros. transitivity a0; auto. +Qed. + +Instance Equal_Equivalence : Equivalence Equal := { + Equivalence_Reflexive := Equal_refl; + Equivalence_Symmetric := Equal_sym; + Equivalence_Transitive := Equal_trans +}. + +Hypothesis eqAdec: EqDec A eqA. + +Program Definition Equal_dec (m1 m2: T.t A) : { m1 === m2 } + { m1 =/= m2 } := + match T.beq (fun a1 a2 => proj_sumbool (a1 == a2)) m1 m2 with + | true => left _ + | false => right _ + end. +Next Obligation. + rename Heq_anonymous into B. + symmetry in B. rewrite T.beq_correct in B. + red; intros. generalize (B x). + destruct (T.get x m1); destruct (T.get x m2); auto. + intros. eapply proj_sumbool_true; eauto. +Qed. +Next Obligation. + assert (T.beq (fun a1 a2 => proj_sumbool (a1 == a2)) m1 m2 = true). + apply T.beq_correct; intros. + generalize (H x). + destruct (T.get x m1); destruct (T.get x m2); try tauto. + intros. apply proj_sumbool_is_true; auto. + unfold equiv, complement in H0. congruence. +Qed. + +Instance Equal_EqDec : EqDec (T.t A) Equal := Equal_dec. + +End EXTENSIONAL_EQUALITY. + +(** Creating a tree from a list of (key, value) pairs. *) + +Section OF_LIST. + +Variable A: Type. + +Let f := fun (m: T.t A) (k_v: T.elt * A) => T.set (fst k_v) (snd k_v) m. + +Definition of_list (l: list (T.elt * A)) : T.t A := + List.fold_left f l (T.empty _). + +Lemma in_of_list: + forall l k v, T.get k (of_list l) = Some v -> In (k, v) l. +Proof. + assert (REC: forall k v l m, + T.get k (fold_left f l m) = Some v -> In (k, v) l \/ T.get k m = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + - tauto. + - apply IHl in H. unfold f in H. simpl in H. rewrite T.gsspec in H. + destruct H; auto. + destruct (T.elt_eq k k1). inv H. auto. auto. + } + intros. apply REC in H. rewrite T.gempty in H. intuition congruence. +Qed. + +Lemma of_list_dom: + forall l k, In k (map fst l) -> exists v, T.get k (of_list l) = Some v. +Proof. + assert (REC: forall k l m, + In k (map fst l) \/ (exists v, T.get k m = Some v) -> + exists v, T.get k (fold_left f l m) = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + - tauto. + - apply IHl. unfold f; rewrite T.gsspec. simpl. destruct (T.elt_eq k k1). + right; econstructor; eauto. + intuition congruence. + } + intros. apply REC. auto. +Qed. + +Remark of_list_unchanged: + forall k l m, ~In k (map fst l) -> T.get k (List.fold_left f l m) = T.get k m. +Proof. + induction l as [ | [k1 v1] l]; simpl; intros. +- auto. +- rewrite IHl by tauto. unfold f; apply T.gso; intuition auto. +Qed. + +Lemma of_list_unique: + forall k v l1 l2, + ~In k (map fst l2) -> T.get k (of_list (l1 ++ (k, v) :: l2)) = Some v. +Proof. + intros. unfold of_list. rewrite fold_left_app. simpl. + rewrite of_list_unchanged by auto. unfold f; apply T.gss. +Qed. + +Lemma of_list_norepet: + forall l k v, list_norepet (map fst l) -> In (k, v) l -> T.get k (of_list l) = Some v. +Proof. + assert (REC: forall k v l m, + list_norepet (map fst l) -> + In (k, v) l -> + T.get k (fold_left f l m) = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + contradiction. + inv H. destruct H0. + inv H. rewrite of_list_unchanged by auto. apply T.gss. + apply IHl; auto. + } + intros; apply REC; auto. +Qed. + +Lemma of_list_elements: + forall m k, T.get k (of_list (T.elements m)) = T.get k m. +Proof. + intros. destruct (T.get k m) as [v|] eqn:M. +- apply of_list_norepet. apply T.elements_keys_norepet. apply T.elements_correct; auto. +- destruct (T.get k (of_list (T.elements m))) as [v|] eqn:M'; auto. + apply in_of_list in M'. apply T.elements_complete in M'. congruence. +Qed. + +End OF_LIST. + +Lemma of_list_related: + forall (A B: Type) (R: A -> B -> Prop) k l1 l2, + list_forall2 (fun ka kb => fst ka = fst kb /\ R (snd ka) (snd kb)) l1 l2 -> + option_rel R (T.get k (of_list l1)) (T.get k (of_list l2)). +Proof. + intros until k. unfold of_list. + set (R' := fun ka kb => fst ka = fst kb /\ R (snd ka) (snd kb)). + set (fa := fun (m : T.t A) (k_v : T.elt * A) => T.set (fst k_v) (snd k_v) m). + set (fb := fun (m : T.t B) (k_v : T.elt * B) => T.set (fst k_v) (snd k_v) m). + assert (REC: forall l1 l2, list_forall2 R' l1 l2 -> + forall m1 m2, option_rel R (T.get k m1) (T.get k m2) -> + option_rel R (T.get k (fold_left fa l1 m1)) (T.get k (fold_left fb l2 m2))). + { induction 1; intros; simpl. + - auto. + - apply IHlist_forall2. unfold fa, fb. rewrite ! T.gsspec. + destruct H as [E F]. rewrite E. destruct (T.elt_eq k (fst b1)). + constructor; auto. + auto. } + intros. apply REC; auto. rewrite ! T.gempty. constructor. +Qed. + +End Tree_Properties. + +Module PTree_Properties := Tree_Properties(PTree). + +(** * Useful notations *) + +Notation "a ! b" := (PTree.get b a) (at level 1). +Notation "a !! b" := (PMap.get b a) (at level 1). diff --git a/src/Predicate.v b/src/Predicate.v new file mode 100644 index 0000000..1ff9831 --- /dev/null +++ b/src/Predicate.v @@ -0,0 +1,1242 @@ +Require Import Coq.Structures.Equalities. + +Require Import TVSMT.Coqlib. +Require Import TVSMT.Maps. +Require Import TVSMT.Errors. +Require Import TVSMT.Sat. +Require Import TVSMT.Common. +Require Import TVSMT.Hashtree. + +Local Open Scope error_monad_scope. +Local Open Scope positive. + +(** The [predicate] type defines the path predicates that are used in the [Ieta] and [Igamma] + instructions. *) +Inductive pred {A}: Type := +| Ptrue : pred +| Pfalse : pred +| Pundef : pred +| Pbase : A -> pred +| Pand : pred -> pred -> pred +| Por : pred -> pred -> pred +| Pimp : pred -> pred -> pred +| Pnot : pred -> pred. + +Module PredicateNotation. + + Declare Scope predicate. + + Notation "A ∧ B" := (Pand A B) (at level 20, left associativity) : predicate. + Notation "A ∨ B" := (Por A B) (at level 25, left associativity) : predicate. + Notation "A → B" := (Pimp A B) (at level 30, right associativity) : predicate. + Notation "¬ A" := (Pnot A) (at level 15) : predicate. + Notation "⟂" := (Pfalse) : predicate. + Notation "'T'" := (Ptrue) : predicate. + Notation "○" := (Pundef) : predicate. + +End PredicateNotation. + +Import PredicateNotation. +Local Open Scope predicate. +Local Open Scope bool. +Local Open Scope positive. + +Module ThreeValued(Base: UsualDecidableTypeBoth). + +Definition A := Base.t. +Definition el_dec: forall a b: A, {a = b} + {a <> b} := Base.eq_dec. + +Module Import AH := HashTree(Base). +Module Import AH_Prop := HashTreeProperties(Base). + +Section EVAL. + +Context (eval: A -> option bool). + +Definition predicate := @pred A. + +(* This evaluation follows Łukasiewicz logic. *) +Fixpoint eval_predicate (p: predicate) : option bool := + match p with + | T => Some true + | ⟂ => Some false + | ○ => None + | ¬ p => + match eval_predicate p with + | None => None + | Some b => Some (negb b) + end + | Pbase c => + match eval c with + | Some p => Some p + | None => None + end + | p1 ∧ p2 => + match eval_predicate p1, eval_predicate p2 with + | Some p1', Some p2' => Some (p1' && p2') + | Some false, None => Some false + | None, Some false => Some false + | _, _ => None + end + | p1 ∨ p2 => + match eval_predicate p1, eval_predicate p2 with + | Some p1', Some p2' => Some (p1' || p2') + | Some true, None => Some true + | None, Some true => Some true + | _, _ => None + end + | p1 → p2 => + match eval_predicate p1, eval_predicate p2 with + | Some p1', Some p2' => Some (negb p1' || p2') + | None, Some false => None + | Some true, None => None + | _, _ => Some true + end + end. + +Definition equiv {A} (a b: @pred A) := + Pand (Pimp a b) (Pimp b a). + +Definition hash_pred := @pred positive. + +Definition hash_tree := PTree.t A. + +Lemma comp_list_dec : + forall (a b: list positive), + {a = b} + {a <> b}. +Proof. generalize Pos.eq_dec. decide equality. Defined. + +Lemma hash_pred_dec : + forall (a b: hash_pred), + {a = b} + {a <> b}. +Proof. generalize Pos.eq_dec. generalize bool_dec. repeat decide equality. Defined. + +Lemma predicate_dec : + forall (a b: predicate), + {a = b} + {a <> b}. +Proof. generalize el_dec. generalize bool_dec. repeat decide equality. Defined. + +Definition find_tree (el: A) (h: hash_tree) : option positive := + match + filter (fun x => match x with (a, b) => if el_dec el b then true else false end) + (PTree.elements h) with + | (p, _) :: nil => Some p + | _ => None + end. + +Definition combine_option {A} (a b: option A) : option A := + match a, b with + | Some a', _ => Some a' + | _, Some b' => Some b' + | _, _ => None + end. + +Definition max_key {A} (t: PTree.t A) := + fold_right Pos.max 1 (map fst (PTree.elements t)). + +Fixpoint hash_predicate (p: predicate) (h: PTree.t A) + : hash_pred * PTree.t A := + match p with + | T => (T, h) + | ⟂ => (⟂, h) + | ○ => (○, h) + | ¬ p => let (p', t) := hash_predicate p h in (¬ p', t) + | Pbase c => + match find_tree c h with + | Some p => (Pbase p, h) + | None => + let nkey := max_key h + 1 in + (Pbase nkey, PTree.set nkey c h) + end + | p1 ∧ p2 => + let (p1', t1) := hash_predicate p1 h in + let (p2', t2) := hash_predicate p2 t1 in + (p1' ∧ p2', t2) + | p1 ∨ p2 => + let (p1', t1) := hash_predicate p1 h in + let (p2', t2) := hash_predicate p2 t1 in + (p1' ∨ p2', t2) + | p1 → p2 => + let (p1', t1) := hash_predicate p1 h in + let (p2', t2) := hash_predicate p2 t1 in + (p1' → p2', t2) + end. + +Fixpoint eval_hash_pred (p: hash_pred) (a: PTree.t bool) : option bool := + match p with + | T => Some true + | ⟂ => Some false + | ○ => None + | ¬ p => + match eval_hash_pred p a with + | Some b => Some (negb b) + | None => None + end + | Pbase k => a ! k + | p1 ∧ p2 => + match eval_hash_pred p1 a, eval_hash_pred p2 a with + | Some p1', Some p2' => Some (p1' && p2') + | Some false, None => Some false + | None, Some false => Some false + | _, _ => None + end + | p1 ∨ p2 => + match eval_hash_pred p1 a, eval_hash_pred p2 a with + | Some p1', Some p2' => Some (p1' || p2') + | Some true, None => Some true + | None, Some true => Some true + | _, _ => None + end + | p1 → p2 => + match eval_hash_pred p1 a, eval_hash_pred p2 a with + | Some p1', Some p2' => Some (negb p1' || p2') + | None, Some false => None + | Some true, None => None + | _, _ => Some true + end + end. + +(** The [predicate] type defines the path predicates that are used in the [Ieta] and [Igamma] + instructions. *) +Inductive pred_bin {A}: Type := +| PBtrue : pred_bin +| PBfalse : pred_bin +| PBbase : (bool * A) -> pred_bin +| PBand : pred_bin -> pred_bin -> pred_bin +| PBor : pred_bin -> pred_bin -> pred_bin. + +Definition hash_pred_bin : Type := @pred_bin positive * @pred_bin positive. + +Fixpoint eval_hash_pred_bin1 (p1: @pred_bin positive) (m: PMap.t bool) : bool := + match p1 with + | PBtrue => true + | PBfalse => false + | PBbase (b, p) => if b then m!!p else negb m!!p + | PBand p1 p2 => eval_hash_pred_bin1 p1 m && eval_hash_pred_bin1 p2 m + | PBor p1 p2 => eval_hash_pred_bin1 p1 m || eval_hash_pred_bin1 p2 m + end. + +Definition eval_hash_pred_bin (p1: hash_pred_bin) (m: PMap.t bool) : (bool * bool) := + (eval_hash_pred_bin1 (fst p1) m, eval_hash_pred_bin1 (snd p1) m). + +Fixpoint negate (h: @pred_bin positive) : @pred_bin positive := + match h with + | PBtrue => PBfalse + | PBfalse => PBtrue + | PBbase (b, p) => PBbase (negb b, p) + | PBand p1 p2 => PBor (negate p1) (negate p2) + | PBor p1 p2 => PBand (negate p1) (negate p2) + end. + +Fixpoint conv_hash_pred (p: hash_pred) (h: PTree.t positive) (fresh: positive): + (hash_pred_bin * PTree.t positive * positive) := + match p with + | Ptrue => ((PBfalse, PBtrue), h, fresh) + | Pfalse => ((PBfalse, PBfalse), h, fresh) + | Pundef => ((PBtrue, PBfalse), h, fresh) + | Pnot p => let '((p1, p2), h', fresh') := conv_hash_pred p h fresh in + ((p1, negate p2), h', fresh') + | Pand p1 p2 => + let '((p1s, p1'), h', fresh') := conv_hash_pred p1 h fresh in + let '((p2s, p2'), h'', fresh'') := conv_hash_pred p2 h' fresh' in + ((PBor (PBand p1' p2s) (PBor (PBand p1s p2s) (PBand p1s p2')), + PBand p1' p2'), h'', fresh'') + | Por p1 p2 => + let '((p1s, p1'), h', fresh') := conv_hash_pred p1 h fresh in + let '((p2s, p2'), h'', fresh'') := conv_hash_pred p2 h' fresh' in + ((PBor (PBand (negate p1') p2s) (PBor (PBand (negate p2') p1s) (PBand p1s p2s)), + PBor p1' p2'), h'', fresh'') + | Pimp p1 p2 => + let '((p1s, p1'), h', fresh') := conv_hash_pred p1 h fresh in + let '((p2s, p2'), h'', fresh'') := conv_hash_pred p2 h' fresh' in + ((PBor (PBand (negate p1s) (PBand p1' p2s)) (PBand (negate p2s) (PBand (negate p2') p1s)), + PBor p1s (PBor (negate p1') p2')), h'', fresh'') + | Pbase k => + ((PBbase (true, fresh), PBbase (true, k)), PTree.set k fresh h, fresh+1) + end. + +Fixpoint mult {A: Type} (a b: list (list A)) : list (list A) := + match a with + | nil => nil + | l :: ls => mult ls b ++ (List.map (fun x => l ++ x) b) + end. + +Definition simplify' {A} (p: @pred A) := + match p with + | A ∧ T => A + | T ∧ A => A + | _ ∧ ⟂ => ⟂ + | ⟂ ∧ _ => ⟂ + | ○ ∧ ○ => ○ + | A ∨ ⟂ => A + | ⟂ ∨ A => A + | _ ∨ T => T + | T ∨ _ => T + | ○ ∨ ○ => ○ + | T → A => A + | ⟂ → _ => T + | _ → T => T + | ○ → ○ => T + | ¬ T => ⟂ + | ¬ ⟂ => T + | ¬ ○ => ○ + | A => A + end. + +Fixpoint simplify {A} (dec: forall a b: @pred A, {a = b} + {a <> b} ) (p: @pred A) := + match p with + | A ∧ B => + let A' := simplify dec A in + let B' := simplify dec B in + if dec A' B' then A' else simplify' (A' ∧ B') + | A ∨ B => + let A' := simplify dec A in + let B' := simplify dec B in + if dec A' B' then A' else simplify' (A' ∨ B') + | A → B => + let A' := simplify dec A in + let B' := simplify dec B in + if dec A' B' then A' else simplify' (A' → B') + | ¬ p => simplify' (¬ (simplify dec p)) + | T => T + | ⟂ => ⟂ + | ○ => ○ + | Pbase a => Pbase a + end. + +Fixpoint sat_hash_pred_bin1' (p: @pred_bin positive) (a: asgn) : bool := + match p with + | PBtrue => true + | PBfalse => false + | PBbase (b, p') => if b then a (Pos.to_nat p') else negb (a (Pos.to_nat p')) + | PBand p1 p2 => sat_hash_pred_bin1' p1 a && sat_hash_pred_bin1' p2 a + | PBor p1 p2 => sat_hash_pred_bin1' p1 a || sat_hash_pred_bin1' p2 a + end. + +Fixpoint sat_hash_pred_bin1 (p: @pred_bin positive) (a: PMap.t bool) : bool := + match p with + | PBtrue => true + | PBfalse => false + | PBbase (b, p') => if b then a !! p' else negb (a !! p') + | PBand p1 p2 => sat_hash_pred_bin1 p1 a && sat_hash_pred_bin1 p2 a + | PBor p1 p2 => sat_hash_pred_bin1 p1 a || sat_hash_pred_bin1 p2 a + end. + +Definition sat_hash_pred_bin (p: hash_pred_bin) (a: PMap.t bool) : bool := + negb (sat_hash_pred_bin1 (fst p) a) && sat_hash_pred_bin1 (snd p) a. + +Definition conv_tree (p: PTree.t positive) (assigns: PTree.t bool) : PMap.t bool := + PTree.fold (fun m i a => + match assigns ! i with + | Some true => PMap.set a false (PMap.set i true m) + | Some false => PMap.set a false (PMap.set i false m) + | None => PMap.set a true (PMap.set i false m) + end) p (PMap.init false). + +Definition interp_bin (b: bool * bool): option bool := + match b with | (true, _) => None | (_, b) => Some b end. + +Fixpoint max_hash_pred (p: hash_pred) := + match p with + | T | ⟂ | ○ => 1 + | Pbase k => k + | p1 ∧ p2 | p1 ∨ p2 | p1 → p2 => Pos.max (max_hash_pred p1) (max_hash_pred p2) + | ¬ p' => max_hash_pred p' + end. + +Definition eval_through_bin (p: hash_pred) (a: PTree.t bool) := + let '(p', h, f) := conv_hash_pred p (PTree.empty _) (max_hash_pred p + 1) in + let m := conv_tree h a in + interp_bin (sat_hash_pred_bin1 (fst p') m, sat_hash_pred_bin1 (snd p') m). + +(*Compute eval_hash_pred (Pand (Por (Pbase 1) (Pbase 2)) (Pbase 3)) + (PTree.set 1 true (PTree.set 3 true (PTree.empty _))). +Compute eval_through_bin (Pand (Por (Pbase 1) (Pbase 2)) (Pbase 3)) + (PTree.set 1 true (PTree.set 3 true (PTree.empty _))).*) + +Lemma satFormula_concat: + forall a b agn, + satFormula a agn -> + satFormula b agn -> + satFormula (a ++ b) agn. +Proof. induction a; crush. Qed. + +Lemma satFormula_concat2: + forall a b agn, + satFormula (a ++ b) agn -> + satFormula a agn /\ satFormula b agn. +Proof. + induction a; simplify; + try apply IHa in H1; crush. +Qed. + +Lemma satClause_concat: + forall a a1 a0, + satClause a a1 -> + satClause (a0 ++ a) a1. +Proof. induction a0; crush. Qed. + +Lemma satClause_concat2: + forall a a1 a0, + satClause a0 a1 -> + satClause (a0 ++ a) a1. +Proof. + induction a0; crush. + inv H; crush. +Qed. + +Lemma satClause_concat3: + forall a b c, + satClause (a ++ b) c -> + satClause a c \/ satClause b c. +Proof. + induction a; crush. + inv H; crush. + apply IHa in H0; crush. + inv H0; crush. +Qed. + +Lemma satFormula_mult': + forall p2 a a0, + satFormula p2 a0 \/ satClause a a0 -> + satFormula (map (fun x : list lit => a ++ x) p2) a0. +Proof. + induction p2; crush. + - inv H. inv H0. apply satClause_concat. auto. + apply satClause_concat2; auto. + - apply IHp2. + inv H; crush; inv H0; crush. +Qed. + +Lemma satFormula_mult2': + forall p2 a a0, + satFormula (map (fun x : list lit => a ++ x) p2) a0 -> + satClause a a0 \/ satFormula p2 a0. +Proof. + induction p2; crush. + apply IHp2 in H1. inv H1; crush. + apply satClause_concat3 in H0. + inv H0; crush. +Qed. + +Lemma satFormula_mult: + forall p1 p2 a, + satFormula p1 a \/ satFormula p2 a -> + satFormula (mult p1 p2) a. +Proof. + induction p1; crush. + apply satFormula_concat; crush. + inv H. inv H0. + apply IHp1. auto. + apply IHp1. auto. + apply satFormula_mult'; + inv H; crush. +Qed. + +Lemma satFormula_mult2: + forall p1 p2 a, + satFormula (mult p1 p2) a -> + satFormula p1 a \/ satFormula p2 a. +Proof. + induction p1; crush. + apply satFormula_concat2 in H; crush. + apply IHp1 in H0. + inv H0; crush. + apply satFormula_mult2' in H1. inv H1; crush. +Qed. + +Fixpoint trans_pred (p: @pred_bin positive) : + {fm: formula | forall a, sat_hash_pred_bin1' p a = true <-> satFormula fm a}. + refine + (match p with + | PBtrue => exist _ nil _ + | PBfalse => exist _ (nil :: nil) _ + | PBbase (b, p') => exist _ (((b, Pos.to_nat p') :: nil) :: nil) _ + | PBand p1 p2 => + match trans_pred p1, trans_pred p2 with + | exist p1' _, exist p2' _ => exist _ (p1' ++ p2') _ + end + | PBor p1 p2 => + match trans_pred p1, trans_pred p2 with + | exist p1' _, exist p2' _ => exist _ (mult p1' p2') _ + end + end); split; crush. + - destruct_match; subst. tauto. simplify. tauto. + - inv H0; crush. destruct_match; crush. subst. apply negb_true_iff. crush. + - apply satFormula_concat. apply i; auto. apply i0; auto. + - apply satFormula_concat2 in H. simplify. apply andb_true_intro. + split. apply i in H0. auto. + apply i0 in H1. auto. + - apply orb_prop in H. inv H; apply satFormula_mult. apply i in H0. auto. + apply i0 in H0. auto. + - apply orb_true_intro. + apply satFormula_mult2 in H. inv H. apply i in H0. auto. + apply i0 in H0. auto. +Qed. + +Definition sat_pred (bound: nat) (p: @pred_bin positive) : + option ({al : alist | sat_hash_pred_bin1' p (interp_alist al) = true} + + {forall a : asgn, sat_hash_pred_bin1' p a = false}). + refine + ( match trans_pred p with + | exist fm _ => + match boundedSat bound fm with + | Some (inleft (exist a _)) => Some (inleft (exist _ a _)) + | Some (inright _) => Some (inright _) + | None => None + end + end ). + - apply i in s1. auto. + - intros. specialize (n a). specialize (i a). + destruct (sat_hash_pred_bin1' p a). exfalso. + apply n. apply i. auto. auto. +Qed. + +Lemma negate_correct : + forall h a, sat_hash_pred_bin1' (negate h) a = negb (sat_hash_pred_bin1' h a). +Proof. + induction h; crush. + - repeat destruct_match; subst; crush; symmetry; apply negb_involutive. + - rewrite negb_andb; crush. + - rewrite negb_orb; crush. +Qed. + +Lemma sat_implication : + forall h h' a, + sat_hash_pred_bin1' (PBand (negate h) h') a = false -> + sat_hash_pred_bin1' h' a = true -> + sat_hash_pred_bin1' h a = true. +Proof. + intros. simplify. rewrite negate_correct in H. + rewrite H0 in H. rewrite andb_true_r in H. crush. +Qed. + +Fixpoint trans_pred_simple (p: @pred_bin positive) := + match p with + | PBtrue => nil + | PBfalse => nil :: nil + | PBbase (b, p') => ((b, Pos.to_nat p') :: nil) :: nil + | PBand p1 p2 => + match trans_pred_simple p1, trans_pred_simple p2 with + | p1', p2' => (p1' ++ p2') + end + | PBor p1 p2 => + match trans_pred_simple p1, trans_pred_simple p2 with + | p1', p2' => (mult p1' p2') + end + end. + +(*Compute trans_pred_simple (PBand (PBbase (false, 1)) (PBbase (true, 3))).*) + +Definition sat_pred_simple (bound: nat) (p: @pred_bin positive) := + boundedSatSimple bound (trans_pred_simple p). + +Definition sat_hash_pred_simple (bound: nat) (p: hash_pred_bin) := + match sat_pred_simple bound (PBor (fst p) (negate (snd p))) with + | Some None => Some None + | Some (Some a) => Some (Some a) + | None => None + end. + +Definition sat_hpred_simple bound (p: hash_pred) := + let '(p', h, f) := conv_hash_pred p (PTree.empty _) (max_hash_pred p + 1) in + sat_hash_pred_simple bound p'. + +Lemma max_key_correct' : + forall l hi, In hi l -> hi <= fold_right Pos.max 1 l. +Proof. + induction l; crush. + inv H. lia. + destruct (Pos.max_dec a (fold_right Pos.max 1 l)); rewrite e. + - apply Pos.max_l_iff in e. + assert (forall a b c, a <= c -> c <= b -> a <= b) by lia. + eapply H; eauto. + - apply IHl; auto. +Qed. + +Definition bar (p1: lit): lit := (negb (fst p1), (snd p1)). + +Definition stseytin_or (cur p1 p2: lit) : formula := + (bar cur :: p1 :: p2 :: nil) + :: (cur :: bar p1 :: nil) + :: (cur :: bar p2 :: nil) :: nil. + +Definition stseytin_and (cur p1 p2: lit) : formula := + (cur :: bar p1 :: bar p2 :: nil) + :: (bar cur :: p1 :: nil) + :: (bar cur :: p2 :: nil) :: nil. + +Fixpoint xtseytin (next: nat) (p: @pred_bin positive) {struct p} : (nat * lit * formula) := + match p with + | PBbase (b, p') => (next, (b, Pos.to_nat p'), nil) + | PBtrue => + ((next+1)%nat, (true, next), ((true, next)::nil)::nil) + | PBfalse => + ((next+1)%nat, (true, next), ((false, next)::nil)::nil) + | PBor p1 p2 => + let '(m1, n1, f1) := xtseytin next p1 in + let '(m2, n2, f2) := xtseytin m1 p2 in + ((m2+1)%nat, (true, m2), stseytin_or (true, m2) n1 n2 ++ f1 ++ f2) + | PBand p1 p2 => + let '(m1, n1, f1) := xtseytin next p1 in + let '(m2, n2, f2) := xtseytin m1 p2 in + ((m2+1)%nat, (true, m2), stseytin_and (true, m2) n1 n2 ++ f1 ++ f2) + end. + +Lemma stseytin_and_correct : + forall cur p1 p2 fm c, + stseytin_and cur p1 p2 = fm -> + satLit cur c -> + satLit p1 c /\ satLit p2 c -> + satFormula fm c. +Proof. + intros. + unfold stseytin_and in *. rewrite <- H. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. +Qed. + +Lemma stseytin_or_correct : + forall cur p1 p2 fm c, + stseytin_or cur p1 p2 = fm -> + satLit cur c -> + satLit p1 c \/ satLit p2 c -> + satFormula fm c. +Proof. + intros. + unfold stseytin_or in *. rewrite <- H. inv H1. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. +Qed. + +Fixpoint max_predicate (p: @pred_bin positive) : positive := + match p with + | PBbase (b, p) => p + | PBtrue => 1 + | PBfalse => 1 + | PBand a b => Pos.max (max_predicate a) (max_predicate b) + | PBor a b => Pos.max (max_predicate a) (max_predicate b) + end. + +Definition tseytin_simple (p: @pred_bin positive) : formula := + let m := Pos.to_nat (max_predicate p + 1) in + let '(m, n, fm) := xtseytin m p in + (n::nil) :: fm. + +Definition sat_pred_simple_tseytin (p: @pred_bin positive) : option (option alist) := + match boundedSatSimple 10000%nat (tseytin_simple p) with + | Some (Some a) => Some (Some a) + | Some None => Some None + | None => None + end. + +Definition sat_hash_pred_simple_tseytin (p: hash_pred_bin) := + match sat_pred_simple_tseytin (PBor (fst p) (negate (snd p))) with + | Some None => Some None + | Some (Some a) => Some (Some a) + | None => None + end. + +Definition sat_hpred_simple_tseytin (p: hash_pred) := + let '(p', h, f) := conv_hash_pred p (PTree.empty _) (max_hash_pred p + 1) in + sat_hash_pred_simple_tseytin p'. + +(*Compute sat_hpred_simple_tseytin ((Pbase 1 ∨ Pbase 2) → (Pbase 1 ∨ Pbase 2)). +Compute sat_hpred_simple 1000%nat (simplify ((Pbase 1 ∨ Pbase 2) → (Pbase 1 ∨ Pbase 2))).*) + +Lemma max_key_correct : + forall A h_tree hi (c: A), + h_tree ! hi = Some c -> + hi <= max_key h_tree. +Proof. + unfold max_key. intros. apply PTree.elements_correct in H. + apply max_key_correct'. + eapply in_map with (f := fst) in H. auto. +Qed. + +Lemma hash_constant : + forall p h h_tree hi c h_tree', + h_tree ! hi = Some c -> + hash_predicate p h_tree = (h, h_tree') -> + h_tree' ! hi = Some c. +Proof. + induction p; crush. + - repeat (destruct_match; crush); subst. + pose proof H as X. apply max_key_correct in X. + rewrite PTree.gso by lia; auto. + - repeat (destruct_match; crush). exploit IHp1; eauto. + - repeat (destruct_match; crush). exploit IHp1; eauto. + - repeat (destruct_match; crush). exploit IHp1; eauto. + - repeat (destruct_match; crush). exploit IHp; eauto. +Qed. + +Lemma find_tree_correct : + forall c h_tree p, + find_tree c h_tree = Some p -> + h_tree ! p = Some c. +Proof. + intros. + unfold find_tree in H. destruct_match; crush. + destruct_match; simplify. + destruct_match; crush. + assert (In (p, a) (filter + (fun x : positive * A => + let (_, b) := x in if el_dec c b then true else false) (PTree.elements h_tree))). + { rewrite Heql. constructor. auto. } + apply filter_In in H. simplify. destruct_match; crush. subst. + apply PTree.elements_complete; auto. +Qed. + +Lemma find_tree_unique : + forall c h_tree p p', + find_tree c h_tree = Some p -> + h_tree ! p' = Some c -> + p = p'. +Proof. + intros. + unfold find_tree in H. + repeat (destruct_match; crush; []). + assert (In (p, a) (filter + (fun x : positive * A => + let (_, b) := x in if el_dec c b then true else false) (PTree.elements h_tree))). + { rewrite Heql. constructor. auto. } + apply filter_In in H. simplify. + destruct (Pos.eq_dec p p'); auto. + exfalso. + destruct_match; subst; crush. + assert (In (p', a) (PTree.elements h_tree) /\ (fun x : positive * A => + let (_, b) := x in if el_dec a b then true else false) (p', a) = true). + { split. apply PTree.elements_correct. auto. rewrite Heqs. auto. } + apply filter_In in H. rewrite Heql in H. inv H. simplify. crush. crush. +Qed. + +Lemma hash_predicate_hash_le : + forall p n p' n', + hash_predicate p n = (p', n') -> + (PTree_Properties.cardinal n <= PTree_Properties.cardinal n')%nat. +Proof. + induction p; crush. + - repeat (destruct_match; crush). + assert (n ! (max_key n + 1) = None). + { destruct (n ! (max_key n + 1)) eqn:?; auto. + apply max_key_correct in Heqo0. lia. } + exploit PTree_Properties.cardinal_set; eauto. + instantiate (1 := a); intros. lia. + - repeat (destruct_match; crush). + apply IHp1 in Heqp. apply IHp2 in Heqp0. lia. + - repeat (destruct_match; crush). + apply IHp1 in Heqp. apply IHp2 in Heqp0. lia. + - repeat (destruct_match; crush). + apply IHp1 in Heqp. apply IHp2 in Heqp0. lia. + - repeat (destruct_match; crush). + apply IHp in Heqp0. lia. +Qed. + +Lemma hash_predicate_hash_len : + forall p n p' n', + hash_predicate p n = (p', n') -> + PTree_Properties.cardinal n = PTree_Properties.cardinal n' -> + n = n'. +Proof. + induction p; crush. + - repeat (destruct_match; crush). + assert (n ! (max_key n + 1) = None). + { destruct (n ! (max_key n + 1)) eqn:?; auto. + apply max_key_correct in Heqo0. lia. } + exploit PTree_Properties.cardinal_set; eauto. + instantiate (1 := a); intros. lia. + - repeat (destruct_match; crush). + pose proof Heqp as X; apply hash_predicate_hash_le in X. + pose proof Heqp0 as X0; apply hash_predicate_hash_le in X0. + exploit IHp1; eauto. lia. intros; subst. + exploit IHp2; eauto. + - repeat (destruct_match; crush). + pose proof Heqp as X; apply hash_predicate_hash_le in X. + pose proof Heqp0 as X0; apply hash_predicate_hash_le in X0. + exploit IHp1; eauto. lia. intros; subst. + exploit IHp2; eauto. + - repeat (destruct_match; crush). + pose proof Heqp as X; apply hash_predicate_hash_le in X. + pose proof Heqp0 as X0; apply hash_predicate_hash_le in X0. + exploit IHp1; eauto. lia. intros; subst. + exploit IHp2; eauto. + - repeat (destruct_match; crush). + pose proof Heqp0 as X; apply hash_predicate_hash_le in X. + exploit IHp; eauto. +Qed. + +Inductive pred_In {A} : A -> @pred A -> Prop := +| pred_In_Pbase : forall a, pred_In a (Pbase a) +| pred_In_Pand1 : forall c p1 p2, pred_In c p1 -> pred_In c (p1 ∧ p2) +| pred_In_Pand2 : forall c p1 p2, pred_In c p2 -> pred_In c (p1 ∧ p2) +| pred_In_Por1 : forall c p1 p2, pred_In c p1 -> pred_In c (p1 ∨ p2) +| pred_In_Por2 : forall c p1 p2, pred_In c p2 -> pred_In c (p1 ∨ p2) +| pred_In_Pimp1 : forall c p1 p2, pred_In c p1 -> pred_In c (p1 → p2) +| pred_In_Pimp2 : forall c p1 p2, pred_In c p2 -> pred_In c (p1 → p2) +| pred_In_Pnot : forall c p, pred_In c p -> pred_In c (¬ p). + +Lemma pred_In_dec : + forall A (c: A) p, + (forall a b : A, { a = b } + { a <> b }) -> + { pred_In c p } + { ~ pred_In c p }. +Proof. + induction p; crush. + - right. unfold not. intros. inv H. + - right. unfold not. intros. inv H. + - right. unfold not. intros. inv H. + - destruct (X c a); subst. + { left. constructor. } + { right. unfold not. intros. inv H. auto. } + - pose proof X as X1. pose proof X as X2. + apply IHp1 in X1. apply IHp2 in X2. + inv X1; inv X2. left. constructor. tauto. + left. constructor. auto. + left. apply pred_In_Pand2. auto. + right. unfold not. intros. inv H1; auto. + - pose proof X as X1. pose proof X as X2. + apply IHp1 in X1. apply IHp2 in X2. + inv X1; inv X2. left. constructor. tauto. + left. constructor. auto. + left. apply pred_In_Por2. auto. + right. unfold not. intros. inv H1; auto. + - pose proof X as X1. pose proof X as X2. + apply IHp1 in X1. apply IHp2 in X2. + inv X1; inv X2. left. constructor. tauto. + left. constructor. auto. + left. apply pred_In_Pimp2. auto. + right. unfold not. intros. inv H1; auto. + - pose proof X as X1. apply IHp in X1. + inv X1. left. constructor. tauto. + right. unfold not. intros. inv H0; auto. +Qed. + +Lemma pred_In_hash : + forall p n p' n' c, + hash_predicate p n = (p', n') -> + pred_In c p -> + exists t, n' ! t = Some c /\ pred_In t p'. +Proof. + induction p; simplify. + - inv H0. + - inv H0. + - inv H0. + - repeat (destruct_match; crush; []); subst. + destruct_match; simplify. + { inv H0. apply find_tree_correct in Heqo. + econstructor; ecrush core. constructor. } + { inv H0. econstructor. rewrite PTree.gss. crush. constructor. } + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Pand2. auto. + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Por2. auto. + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Pimp2. auto. + - repeat (destruct_match; crush; []). inv H0. exploit IHp; eauto; simplify. + econstructor; simplify; eauto. econstructor; auto. +Qed. + +Lemma hash_pred_In' : + forall p h t c n n', + hash_predicate p n = (h, n') -> + n ! t = None -> + n' ! t = Some c -> + pred_In c p. +Proof. + induction p; crush. + - repeat (destruct_match); subst; crush. + destruct (Pos.eq_dec t (max_key n + 1)); subst. + { rewrite PTree.gss in H1. simplify. constructor. } + { rewrite PTree.gso in H1; crush. } + - repeat (destruct_match); subst; crush. + destruct (t0 ! t) eqn:?. + { pose proof Heqo as X. eapply hash_constant in X; eauto. simplify. + constructor. eapply IHp1; eauto. } + { apply pred_In_Pand2; eapply IHp2; eauto. } + - repeat (destruct_match); subst; crush. + destruct (t0 ! t) eqn:?. + { pose proof Heqo as X. eapply hash_constant in X; eauto. simplify. + constructor. eapply IHp1; eauto. } + { apply pred_In_Por2; eapply IHp2; eauto. } + - repeat (destruct_match); subst; crush. + destruct (t0 ! t) eqn:?. + { pose proof Heqo as X. eapply hash_constant in X; eauto. simplify. + constructor. eapply IHp1; eauto. } + { apply pred_In_Pimp2; eapply IHp2; eauto. } + - repeat (destruct_match); subst; crush. + exploit IHp; eauto; intros. econstructor; auto. +Qed. + +Lemma hash_pred_In : + forall p h p' h' c n, + hash_predicate p (PTree.empty _) = (h, n) -> + hash_predicate p' n = (h', n) -> + pred_In c p' -> pred_In c p. +Proof. + intros. exploit pred_In_hash; eauto. simplify. + eapply hash_pred_In'; eauto. apply PTree.gempty. +Qed. + +Lemma pred_In_hash2 : + forall p h n n' t, + hash_predicate p n = (h, n') -> + pred_In t h -> + exists c, n' ! t = Some c /\ pred_In c p. +Proof. + induction p; simplify. + - inv H0. + - inv H0. + - inv H0. + - repeat (destruct_match; crush; []); subst. + destruct_match; simplify. + { inv H0. apply find_tree_correct in Heqo. + econstructor; ecrush core. constructor. } + { inv H0. econstructor. rewrite PTree.gss. crush. constructor. } + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Pand2. auto. + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Por2. auto. + - repeat (destruct_match; crush; []). inv H0. + exploit IHp1; eauto. simplify. econstructor. simplify. eapply hash_constant; eauto. + constructor; auto. + exploit IHp2; eauto. simplify. econstructor; ecrush core. apply pred_In_Pimp2. auto. + - repeat (destruct_match; crush; []). inv H0. + exploit IHp; eauto; simplify; repeat econstructor; ecrush core. +Qed. + +Lemma hash_constant2 : + forall p n h n' v c c', + hash_predicate p n = (h, n') -> + n ! v = Some c -> + n' ! v = Some c' -> + c = c'. +Proof. + intros. + eapply hash_constant in H; eauto. rewrite H1 in H. crush. +Qed. + +Lemma pred_not_in_Pand : + forall A (c: A) p1 p2, + ~ pred_In c (p1 ∧ p2) -> + ~ pred_In c p1 /\ ~ pred_In c p2. +Proof. + intros. split. + - unfold not in *. intros. apply H. constructor. auto. + - unfold not in *. intros. apply H. apply pred_In_Pand2. auto. +Qed. + +Lemma pred_not_in_Por : + forall A (c: A) p1 p2, + ~ pred_In c (p1 ∨ p2) -> + ~ pred_In c p1 /\ ~ pred_In c p2. +Proof. + intros. split. + - unfold not in *. intros. apply H. constructor. auto. + - unfold not in *. intros. apply H. apply pred_In_Por2. auto. +Qed. + +Lemma simplify'_correct : + forall h a, + eval_hash_pred (simplify' h) a = eval_hash_pred h a. +Proof. destruct h; crush; repeat (destruct_match; crush). Qed. + +Lemma eval_hash_pred_Pand_destr : + forall p1 p2 p1' p2' a, + eval_hash_pred p1 a = eval_hash_pred p1' a -> + eval_hash_pred p2 a = eval_hash_pred p2' a -> + eval_hash_pred (Pand p1 p2) a = eval_hash_pred (Pand p1' p2') a. +Proof. intros. simpl. rewrite H in *. rewrite H0 in *. auto. Qed. + +Lemma eval_hash_pred_Pand_symm : + forall p1 p2 a, + eval_hash_pred (Pand p1 p2) a = eval_hash_pred (Pand p2 p1) a. +Proof. simplify. repeat destruct_match; auto. Qed. + +Lemma eval_hash_pred_Por_symm : + forall p1 p2 a, + eval_hash_pred (Por p1 p2) a = eval_hash_pred (Por p2 p1) a. +Proof. simplify. repeat destruct_match; auto. Qed. + +Lemma eval_hash_pred_Por_destr : + forall p1 p2 p1' p2' a, + eval_hash_pred p1 a = eval_hash_pred p1' a -> + eval_hash_pred p2 a = eval_hash_pred p2' a -> + eval_hash_pred (Por p1 p2) a = eval_hash_pred (Por p1' p2') a. +Proof. intros. simpl. rewrite H in *. rewrite H0 in *. auto. Qed. + +Lemma eval_hash_pred_Pand_refl : + forall p1 a, + eval_hash_pred (Pand p1 p1) a = eval_hash_pred p1 a. +Proof. intros. simpl. repeat destruct_match; auto. Qed. + +Lemma eval_hash_pred_Por_refl : + forall p1 a, + eval_hash_pred (Por p1 p1) a = eval_hash_pred p1 a. +Proof. intros. simpl. repeat destruct_match; auto. Qed. + + Lemma pand_false : + forall a b l, + eval_hash_pred (Pand a b) l <> Some true -> + eval_hash_pred b l = Some true -> + eval_hash_pred a l <> Some true. +Proof. simplify. repeat destruct_match; crush. Qed. + +Lemma pand_true : + forall a b l, + eval_hash_pred (Pand a b) l = Some true -> + eval_hash_pred a l = Some true /\ eval_hash_pred b l = Some true. +Proof. simplify; repeat destruct_match; crush. Qed. + +Lemma pand_true2 : + forall a b l, + eval_hash_pred b l = Some true -> + eval_hash_pred a l = Some true -> + eval_hash_pred (Pand a b) l = Some true. +Proof. simplify; repeat destruct_match; crush. Qed. + +Lemma eval_hash_pred_fold_Pand : + forall l x la, + eval_hash_pred (fold_left Pand l x) la = Some true -> + eval_hash_pred x la = Some true. +Proof. + induction l as [| el l' IHl ]; [tauto|]. + intros * EVAL. cbn in *. + eapply IHl in EVAL. eapply pand_true in EVAL; tauto. +Qed. + +Lemma eval_hash_pred_fold_Pand2 : + forall l x la y, + eval_hash_pred (fold_left Pand l x) la = Some true -> + eval_hash_pred y la = Some true -> + eval_hash_pred (fold_left Pand l y) la = Some true. +Proof. + induction l as [| el l' IHl ]; [tauto|]. + intros * EVAL EVALY. cbn in *. + pose proof EVAL as EVAL2. + apply eval_hash_pred_fold_Pand in EVAL. + eapply IHl; eauto. apply pand_true2; apply pand_true in EVAL; tauto. +Qed. + +Lemma eval_hash_pred_fold_Pand3 : + forall l x la, + eval_hash_pred (fold_left Pand l x) la = Some true -> + eval_hash_pred (fold_left Pand l T) la = Some true. +Proof. eauto using eval_hash_pred_fold_Pand2. Qed. + +Lemma eval_hash_pred_fold_Pand4 : + forall l x la, + eval_hash_pred (fold_left Pand l x) la = Some true -> + eval_hash_pred (fold_left Pand l T) la = Some true /\ eval_hash_pred x la = Some true. +Proof. + intros; split; + eauto using eval_hash_pred_fold_Pand3, eval_hash_pred_fold_Pand. +Qed. + +Lemma eval_equiv2 : + forall a b la, + eval_hash_pred a la = eval_hash_pred b la -> + eval_hash_pred (equiv a b) la = Some true. +Proof. + unfold equiv; intros. unfold eval_hash_pred in *. fold eval_hash_pred in *. + repeat destruct_match; try discriminate; crush. +Qed. + +Lemma eval_equiv : + forall a b la, + eval_hash_pred (equiv a b) la = Some true -> + eval_hash_pred a la = eval_hash_pred b la. +Proof. + unfold equiv; intros. unfold eval_hash_pred in *. fold eval_hash_pred in *. + repeat destruct_match; try discriminate; crush. +Qed. + +Lemma eval_hash_pred_T_Pand : + forall a l, + eval_hash_pred (Pand T a) l = eval_hash_pred a l. +Proof. + intros; unfold eval_hash_pred; fold eval_hash_pred; + destruct_match; auto. +Qed. + +Lemma fold_and_forall : + forall p a pt, + eval_hash_pred (fold_left Pand p pt) a = Some true -> + Forall (fun x => eval_hash_pred x a = Some true) p. +Proof. + induction p; crush. + constructor. + - apply eval_hash_pred_fold_Pand in H. eapply pand_true; eauto. + - eapply IHp. eauto. +Qed. + +Lemma fold_and_forall2 : + forall p a pt, + eval_hash_pred pt a = Some true -> + Forall (fun x => eval_hash_pred x a = Some true) p -> + eval_hash_pred (fold_left Pand p pt) a = Some true. +Proof. + induction p; crush. inv H0. + eapply IHp; eauto. + apply pand_true2; eauto. +Qed. + +Lemma fold_and_forall3 : + forall p a, + Forall (fun x => eval_hash_pred x a = Some true) p <-> + eval_hash_pred (fold_left Pand p T) a = Some true. +Proof. + intros; split; intros. + eapply fold_and_forall2; eauto. + eapply fold_and_forall; eauto. +Qed. + +Lemma eval_hash_pred_gso2 : + forall p x y b, + ~ pred_In y p -> + eval_hash_pred p (PTree.set y b x) = eval_hash_pred p x. +Proof. + induction p; try solve [crush]. + - intros * MAX. cbn in *. rewrite PTree.gso; auto. unfold not; intros. + apply MAX; subst; constructor. + - intros. cbn in H. simplify. erewrite !IHp1. + erewrite !IHp2. auto. unfold not in *; intros; apply H; now apply pred_In_Pand2. + unfold not in *; intros; apply H; now constructor. + - intros. cbn in H. simplify. erewrite !IHp1. + erewrite !IHp2. auto. unfold not in *; intros; apply H; now apply pred_In_Por2. + unfold not in *; intros; apply H; now constructor. + - intros. cbn in H. simplify. erewrite !IHp1. + erewrite !IHp2. auto. unfold not in *; intros; apply H; now apply pred_In_Pimp2. + unfold not in *; intros; apply H; now constructor. + - intros. cbn in H. simplify. erewrite !IHp. auto. + unfold not in *; intros; apply H; now constructor. +Qed. + +Lemma eval_hash_pred_gso : + forall p x y b, + (max_hash_pred p < y)%positive -> + eval_hash_pred p (PTree.set y b x) = eval_hash_pred p x. +Proof. + induction p; try solve [crush]. + - intros * MAX. cbn in *. now rewrite PTree.gso by lia. + - intros. cbn in H. simplify. erewrite !IHp1 by lia. + now erewrite !IHp2 by lia. + - intros. cbn in H. simplify. erewrite !IHp1 by lia. + now erewrite !IHp2 by lia. + - intros. cbn in H. simplify. erewrite !IHp1 by lia. + now erewrite !IHp2 by lia. + - intros. cbn in H. simplify. now erewrite !IHp by lia. +Qed. + +Lemma eval_hash_pred_get : + forall p a, + eval_hash_pred (Pbase p) a = a ! p. +Proof. crush. Qed. + +Lemma eval_hash_pred_except : + forall p a a0 a0', + ~ pred_In a p -> + (forall x : positive, x <> a -> a0 ! x = a0' ! x) -> + eval_hash_pred p a0 = eval_hash_pred p a0'. +Proof. + induction p; crush. + - destruct (peq a0 a); subst. + + exfalso. apply H. constructor. + + now rewrite H0 by auto. + - erewrite IHp1; eauto. + erewrite IHp2; eauto. + unfold not; intros; apply H. apply pred_In_Pand2; auto. + unfold not; intros; apply H. constructor; auto. + - erewrite IHp1; eauto. + erewrite IHp2; eauto. + unfold not; intros; apply H. apply pred_In_Por2; auto. + unfold not; intros; apply H. constructor; auto. + - erewrite IHp1; eauto. + erewrite IHp2; eauto. + unfold not; intros; apply H. apply pred_In_Pimp2; auto. + unfold not; intros; apply H. constructor; auto. + - erewrite IHp; eauto. + unfold not; intros; apply H. constructor; auto. +Qed. + +Lemma eval_hash_pred_pand : + forall p p1 p2 p3 a a', + eval_hash_pred p1 a = eval_hash_pred p a' -> + eval_hash_pred p2 a = eval_hash_pred p3 a' -> + eval_hash_pred (p1 ∧ p2) a = eval_hash_pred (p ∧ p3) a'. +Proof. simplify; rewrite !H; now rewrite !H0. Qed. + +Lemma eval_hash_pred_por : + forall p p1 p2 p3 a a', + eval_hash_pred p1 a = eval_hash_pred p a' -> + eval_hash_pred p2 a = eval_hash_pred p3 a' -> + eval_hash_pred (p1 ∨ p2) a = eval_hash_pred (p ∨ p3) a'. +Proof. simplify; rewrite !H; now rewrite !H0. Qed. + +Lemma eval_hash_pred_pimp : + forall p p1 p2 p3 a a', + eval_hash_pred p1 a = eval_hash_pred p a' -> + eval_hash_pred p2 a = eval_hash_pred p3 a' -> + eval_hash_pred (p1 → p2) a = eval_hash_pred (p → p3) a'. +Proof. simplify; rewrite !H; now rewrite !H0. Qed. + +Lemma eval_hash_pred_pnot : + forall p p' a a', + eval_hash_pred p a = eval_hash_pred p' a' -> + eval_hash_pred (¬ p) a = eval_hash_pred (¬ p') a'. +Proof. simplify; now rewrite !H. Qed. + +Lemma fold_left_Pand_rev : + forall l b, + eval_hash_pred (fold_left Pand (rev l) T) b = Some true -> + eval_hash_pred (fold_left Pand l T) b = Some true. +Proof. + intros. + apply fold_and_forall3. + apply fold_and_forall3 in H. + apply Forall_forall; intros. + now eapply Forall_forall in H; [|rewrite <- in_rev; eassumption]. +Qed. + +Lemma fold_left_Pand_rev2 : + forall l b, + eval_hash_pred (fold_left Pand l T) b = Some true -> + eval_hash_pred (fold_left Pand (rev l) T) b = Some true. +Proof. + intros. + apply fold_and_forall3. + apply fold_and_forall3 in H. + apply Forall_forall; intros. + now eapply Forall_forall in H; [|rewrite in_rev; eassumption]. +Qed. + +Lemma eval_hash_pred_T_Pand2 : + forall a b l, + eval_hash_pred b l = Some true -> + eval_hash_pred (Pand a b) l = eval_hash_pred a l. +Proof. + intros; unfold eval_hash_pred; fold eval_hash_pred. + repeat (destruct_match; crush). +Qed. + +Lemma eval_hash_pred_T_Pand3: + forall b a, + eval_hash_pred (b ∧ T) a = eval_hash_pred b a. +Proof. + unfold eval_hash_pred; crush. + repeat (destruct_match; crush). +Qed. + +End EVAL. + +End ThreeValued. diff --git a/src/Sat.v b/src/Sat.v new file mode 100644 index 0000000..1ca3868 --- /dev/null +++ b/src/Sat.v @@ -0,0 +1,580 @@ +(** Homework Assignment 6#
# +##Interactive Computer Theorem +Proving#
# +CS294-9, Fall 2006#
# +UC Berkeley *) + +(** Submit your solution file for this assignment as an attachment + ##by e-mail## with + the subject "ICTP HW6" by the start of class on October 12. + You should write your solutions entirely on your own, which includes not + consulting any solutions to these problems that may be posted on the web. + + ##Template source file## + *) + +Require Import Arith Bool List. + +(** This assignment involves building a certified boolean satisfiability solver + based on the DPLL algorithm. Your certified procedure will take as input a + boolean formula in conjunctive normal form (CNF) and either return a + satisfying assignment to the variables or a value signifying that the input + formula is unsatisfiable. Moreover, the procedure will be implemented with a + rich specification, so you'll know that the answer it gives is correct. By + the end of the assignment, you'll have extracted OCaml code that can be used + to solve some of the more modest classes of problems in the SATLIB archive. + + If you need to page in the relevant background material, try the Wikipedia + pages on + ##SAT## + and + ##the DPLL + algorithm##. The implementation we'll develop here omits the pure literal + heuristic mentioned on the Wikipedia page but is otherwise identical. + *) + + +(** * Problem Definition *) + +Definition var := nat. +(** We identify propositional variables with natural numbers. *) + +Definition lit := (bool * var)%type. +(** A literal is a combination of a truth value and a variable. *) + +Definition clause := list lit. +(** A clause is a list (disjunction) of literals. *) + +Definition formula := list clause. +(** A formula is a list (conjunction) of clauses. *) + +Definition asgn := var -> bool. +(** An assignment is a function from variables to their truth values. *) + +Definition satLit (l : lit) (a : asgn) := + a (snd l) = fst l. +(** An assignment satisfies a literal if it maps it to true. *) + +Fixpoint satClause (cl : clause) (a : asgn) {struct cl} : Prop := + match cl with + | nil => False + | l :: cl' => satLit l a \/ satClause cl' a + end. +(** An assignment satisfies a clause if it satisfies at least one of its + literals. + *) + +Fixpoint satFormula (fm: formula) (a: asgn) {struct fm} : Prop := + match fm with + | nil => True + | cl :: fm' => satClause cl a /\ satFormula fm' a + end. +(** An assignment satisfies a formula if it satisfies all of its clauses. *) + +(** * Subroutines *) + +(** This is the only section of this assignment where you need to provide your + own solutions. You will be implementing four crucial subroutines used by + DPLL. + + I've provided a number of useful definitions and lemmas which you should feel + free to take advantage of in your definitions. A few tips to keep in mind + while writing these strongly specified functions: + - You have a case-by-case choice of a "programming" approach, based around the + [refine] tactic; or a "proving" approach, where the "code" parts of your + definitions are constructed step by step with tactics. The former is often + harder to get started with, but it tends to be more maintainable. + - When you use [refine] with a [fix] expression, it's usually a good idea to + use the [clear] tactic to remove the recursive function name from the + context immediately afterward. This is because Coq won't check that you + call this function with strictly smaller arguments until the whole proof is + done, and it's a real downer to be told you had an invalid recursion + somewhere after you finally "finish" a proof. Instead, make all recursive + calls explicit in the [refine] argument and clear the function name + afterward. + - You'll probably end up with a lot of proof obligations to discharge, and you + definitely won't want to prove most of them manually. These tactics will + probably be your best friends here: [intuition], [firstorder], [eauto], + [simpl], [subst], .... You will probably want to follow your [refine] tactics + with semicolons and strings of semicolon-separated tactics. These strings + should probably start out with basic simplifiers like [intros], [simpl], and + [subst]. + - A word of warning about the [firstorder] tactic: When it works, it works + really well! However, this tactic has a way of running forever on + complicated enough goals. Be ready to cancel its use (e.g., press the + "Stop" button in Proof General) if it takes more than a few seconds. If + you do things the way I have, be prepared to mix and match all sorts of + different combinations of the automating tactics to get a proof script that + solves the problem quickly enough. + - The dependent type families that we use with rich specifications are all + defined in ##the + Specif module## of the Coq standard library. One potential gotcha when + using them comes from the fact that they are defined inductively with + parameters; that is, some arguments to these type families are defined + before the colon in the [Inductive] command. Compared to general arguments + stemming from function types after that colon, usage of parameters is + restricted; they aren't allowed to vary in recursive occurrences of the + type being defined, for instance. Because of this, parameters are ignored + for the purposes of pattern-matching, while they must be passed when + actually constructing new values. For instance, one would pattern-match a + value of a [sig] type with a pattern like [exist x pf], while one would + construct a new value of the same type like [exist _ x pf]. The parameter + [P] is passed in the second case, and we use an underscore when the Coq + type-checker ought to be able to infer its value. When this inference isn't + possible, you may need to specify manually the predicate defining the [sig] + type you want. + + You can also consult the sizeable example at the end of this file, which ties + together the pieces you are supposed to write. + *) + +(** You'll probably want to compare booleans for equality at some point. *) +Definition bool_eq_dec : forall (x y : bool), {x = y} + {x <> y}. + decide equality. +Defined. + +(** A literal and its negation can't be true simultaneously. *) +Lemma contradictory_assignment : forall s l cl a, + s <> fst l + -> satLit l a + -> satLit (s, snd l) a + -> satClause cl a. + intros. + red in H0, H1. + simpl in H1. + subst. + tauto. +Qed. + +Local Hint Resolve contradictory_assignment : core. + +(** Augment an assignment with a new mapping. *) +Definition upd (a : asgn) (l : lit) : asgn := + fun v : var => + if eq_nat_dec v (snd l) + then fst l + else a v. + +(** Some lemmas about [upd] *) + +Lemma satLit_upd_eq : forall l a, + satLit l (upd a l). + unfold satLit, upd; simpl; intros. + destruct (eq_nat_dec (snd l) (snd l)); tauto. +Qed. + +Local Hint Resolve satLit_upd_eq : core. + +Lemma satLit_upd_neq : forall v l s a, + v <> snd l + -> satLit (s, v) (upd a l) + -> satLit (s, v) a. + unfold satLit, upd; simpl; intros. + destruct (eq_nat_dec v (snd l)); tauto. +Qed. + +Local Hint Resolve satLit_upd_neq : core. + +Lemma satLit_upd_neq2 : forall v l s a, + v <> snd l + -> satLit (s, v) a + -> satLit (s, v) (upd a l). + unfold satLit, upd; simpl; intros. + destruct (eq_nat_dec v (snd l)); tauto. +Qed. + +Local Hint Resolve satLit_upd_neq2 : core. + +Lemma satLit_contra : forall s l a cl, + s <> fst l + -> satLit (s, snd l) (upd a l) + -> satClause cl a. + unfold satLit, upd; simpl; intros. + destruct (eq_nat_dec (snd l) (snd l)); intuition. + assert False; intuition. +Qed. + +Local Hint Resolve satLit_contra : core. + +(** Here's the tactic that I used to discharge ##all## proof obligations + in my implementations of the four functions you will define. + It comes with no warranty, as different implementations may lead to + obligations that it can't solve, or obligations that it takes 42 years to + solve. + However, if you think enough like me, each of the four definitions you fill in + should read like: [[ +refine some_expression_with_holes; clear function_name; magic_solver. +]] leaving out the [clear] invocation for non-recursive function definitions. + *) +Ltac magic_solver := simpl; intros; subst; intuition eauto; firstorder; + match goal with + | [ H1 : satLit ?l ?a, H2 : satClause ?cl ?a |- _ ] => + assert (satClause cl (upd a l)); firstorder + end. + +(** OK, here's your first challenge. Write this strongly-specified function to + update a clause to reflect the effect of making a particular literal true. + *) +Definition setClause : forall (cl : clause) (l : lit), + {cl' : clause | + forall a, satClause cl (upd a l) <-> satClause cl' a} + + {forall a, satLit l a -> satClause cl a}. + refine (fix setClause (cl: clause) (l: lit) {struct cl} := + match cl with + | nil => inleft (exist _ nil _) + | x :: xs => + match setClause xs l with + | inright p => inright _ + | inleft (exist _ cl' H) => + match eq_nat_dec (snd x) (snd l), bool_eq_dec (fst x) (fst l) with + | left nat_eq, left bool_eq => inright _ + | left eq, right ne => inleft (exist _ cl' _) + | right neq, _ => inleft (exist _ (x :: cl') _) + end + end + end); clear setClause; try magic_solver. + - intros; simpl; left; apply injective_projections in bool_eq; subst; auto. + - split; intros. simpl in H0. inversion H0. eapply satLit_contra; eauto. + destruct x; simpl in *; subst. auto. + apply H. eauto. + simpl. right. apply H; eauto. + - split; intros; simpl in *. inversion H0. destruct x; simpl in *; subst. + left. eauto. + right. apply H. eauto. + inversion H0. left. apply satLit_upd_neq2. auto. auto. + right. apply H. auto. +Defined. + +(** For testing purposes, we define a weakly-specified function as a thin + wrapper around [setClause]. + *) +Definition setClauseSimple (cl : clause) (l : lit) := + match setClause cl l with + | inleft (exist _ cl' _) => Some cl' + | inright _ => None + end. + +(** When your [setClause] implementation is done, you should be able to + uncomment these test cases and verify that each one yields the correct answer. + Be sure that your [setClause] definition ends in [Defined] and not [Qed], as + the former exposes the definition for use in computational reduction, while + the latter doesn't. + *) + +(*Eval compute in setClauseSimple ((false, 1%nat) :: nil) (true, 1%nat).*) +(*Eval compute in setClauseSimple nil (true, 0). +Eval compute in setClauseSimple ((true, 0) :: nil) (true, 0). +Eval compute in setClauseSimple ((true, 0) :: nil) (false, 0). +Eval compute in setClauseSimple ((true, 0) :: nil) (true, 1). +Eval compute in setClauseSimple ((true, 0) :: nil) (false, 1). +Eval compute in setClauseSimple ((true, 0) :: (true, 1) :: nil) (true, 1). +Eval compute in setClauseSimple ((true, 0) :: (true, 1) :: nil) (false, 1). +Eval compute in setClauseSimple ((true, 0) :: (false, 1) :: nil) (true, 1). +Eval compute in setClauseSimple ((true, 0) :: (false, 1) :: nil) (false, 1).*) + +(** It's useful to have this strongly-specified nilness check. *) +Definition isNil : forall (A : Set) (ls : list A), {ls = nil} + {True}. + destruct ls; [left|right]; eauto. +Defined. +Arguments isNil [A]. + +(** Some more lemmas that I found helpful.... *) + +Lemma satLit_idem_lit : forall l a l', + satLit l a + -> satLit l' a + -> satLit l' (upd a l). + unfold satLit, upd; simpl; intros. + destruct (eq_nat_dec (snd l') (snd l)); congruence. +Qed. + +Local Hint Resolve satLit_idem_lit : core. + +Lemma satLit_idem_clause : forall l a cl, + satLit l a + -> satClause cl a + -> satClause cl (upd a l). + induction cl; simpl; intuition. +Qed. + +Local Hint Resolve satLit_idem_clause : core. + +Lemma satLit_idem_formula : forall l a fm, + satLit l a + -> satFormula fm a + -> satFormula fm (upd a l). + induction fm; simpl; intuition. +Qed. + +Local Hint Resolve satLit_idem_formula : core. + +(** Challenge 2: Write this function that updates an entire formula to reflect + setting a literal to true. + *) +Definition setFormula : forall (fm : formula) (l : lit), + {fm' : formula | + forall a, satFormula fm (upd a l) <-> satFormula fm' a} + + {forall a, satLit l a -> ~satFormula fm a}. + refine (fix setFormula (fm: formula) (l: lit) {struct fm} := + match fm with + | nil => inleft (exist _ nil _) + | cl' :: fm' => + match setFormula fm' l with + | inright p => inright _ + | inleft (exist _ fm'' H) => + match setClause cl' l with + | inright H' => inleft (exist _ fm'' _) + | inleft (exist _ cl'' _) => + if isNil cl'' + then inright _ + else inleft (exist _ (cl'' :: fm'') _) + end + end + end); clear setFormula; try solve [magic_solver]. +Defined. + +(** Here's some code for testing your implementation. *) + +Definition setFormulaSimple (fm : formula) (l : lit) := + match setFormula fm l with + | inleft (exist _ fm' _) => Some fm' + | inright _ => None + end. + +(*Eval compute in setFormulaSimple (((true, 1%nat) :: nil) :: ((false, 1%nat) :: nil) :: nil) (true, 1%nat). +Eval compute in setFormulaSimple nil (true, 0). +Eval compute in setFormulaSimple (((true, 0) :: nil) :: nil) (true, 0). +Eval compute in setFormulaSimple (((false, 0) :: nil) :: nil) (true, 0). +Eval compute in setFormulaSimple (((false, 1) :: nil) :: nil) (true, 0). +Eval compute in setFormulaSimple (((false, 1) :: (true, 0) :: nil) :: nil) (true, 0). +Eval compute in setFormulaSimple (((false, 1) :: (false, 0) :: nil) :: nil) (true, 0).*) + +Local Hint Extern 1 False => discriminate : core. + +Local Hint Extern 1 False => + match goal with + | [ H : In _ (_ :: _) |- _ ] => inversion H; clear H; subst + end : core. + +(** Challenge 3: Write this code that either finds a unit clause in a formula + or declares that there are none. + *) +Definition findUnitClause : forall (fm: formula), + {l : lit | In (l :: nil) fm} + + {forall l, ~In (l :: nil) fm}. + refine (fix findUnitClause (fm: formula) {struct fm} := + match fm with + | nil => inright _ + | (l :: nil) :: fm' => inleft (exist _ l _) + | cl :: fm' => + match findUnitClause fm' with + | inleft (exist _ l _) => inleft (exist _ l _) + | inright H => inright _ + end + end + ); clear findUnitClause; magic_solver. +Defined. + +(** The literal in a unit clause must always be true in a satisfying + assignment. + *) +Lemma unitClauseTrue : forall l a fm, + In (l :: nil) fm + -> satFormula fm a + -> satLit l a. + induction fm; intuition. + inversion H. + inversion H; subst; simpl in H0; intuition. +Qed. + +Local Hint Resolve unitClauseTrue : core. + +(** Final challenge: Implement unit propagation. The return type of + [unitPropagate] signifies that three results are possible: + - [None]: There are no unit clauses. + - [Some (inleft _)]: There is a unit clause, and here is a formula reflecting + setting its literal to true. + - [Some (inright _)]: There is a unit clause, and propagating it reveals that + the formula is unsatisfiable. + *) +Definition unitPropagate : forall (fm : formula), option (sigT (fun fm' : formula => + {l : lit | + (forall a, satFormula fm a -> satLit l a) + /\ forall a, satFormula fm (upd a l) <-> satFormula fm' a}) ++ {forall a, ~satFormula fm a}). + refine (fix unitPropagate (fm: formula) {struct fm} := + match findUnitClause fm with + | inright H => None + | inleft (exist _ l _) => + match setFormula fm l with + | inright _ => Some (inright _) + | inleft (exist _ fm' _) => + Some (inleft (existT _ fm' (exist _ l _))) + end + end + ); clear unitPropagate; magic_solver. +Defined. + + +Definition unitPropagateSimple (fm : formula) := + match unitPropagate fm with + | None => None + | Some (inleft (existT _ fm' (exist _ l _))) => Some (Some (fm', l)) + | Some (inright _) => Some None + end. + +(*Eval compute in unitPropagateSimple (((true, 1%nat) :: nil) :: ((false, 1%nat) :: nil) :: nil). +Eval compute in unitPropagateSimple nil. +Eval compute in unitPropagateSimple (((true, 0) :: nil) :: nil). +Eval compute in unitPropagateSimple (((true, 0) :: nil) :: ((false, 0) :: nil) :: nil). +Eval compute in unitPropagateSimple (((true, 0) :: nil) :: ((false, 1) :: nil) :: nil). +Eval compute in unitPropagateSimple (((true, 0) :: nil) :: ((false, 0) :: (false, 1) :: nil) :: nil). +Eval compute in unitPropagateSimple (((false, 0) :: (false, 1) :: nil) :: ((true, 0) :: nil) :: nil).*) + + +(** * The SAT Solver *) + +(** This section defines a DPLL SAT solver in terms of the subroutines you've + written. + *) + +(** An arbitrary heuristic to choose the next variable to split on *) +Definition chooseSplit (fm : formula) := + match fm with + | ((l :: _) :: _) => l + | _ => (true, 0) + end. + +Definition negate (l : lit) := (negb (fst l), snd l). + +Local Hint Unfold satFormula : core. + +Lemma satLit_dec : forall l a, + {satLit l a} + {satLit (negate l) a}. + destruct l. + unfold satLit; simpl; intro. + destruct b; destruct (a v); simpl; auto. +Qed. + +(** We'll represent assignments as lists of literals instead of functions. *) +Definition alist := list lit. + +Fixpoint interp_alist (al : alist) : asgn := + match al with + | nil => fun _ => true + | l :: al' => upd (interp_alist al') l + end. + +(** Here's the final definition! This is not the way you should write proof + scripts. ;-) + + This implementation isn't ##quite## what you would expect, since it + takes an extra parameter expressing a search tree depth. Writing the function + without that parameter would be trickier when it came to proving termination. + In practice, you can just seed the bound with one plus the number of variables + in the input, but the function's return type still indicates a possibility for + a "time-out" by returning [None]. + *) +Definition boundedSat: forall (bound : nat) (fm : formula), + option ({al : alist | satFormula fm (interp_alist al)} + + {forall a, ~satFormula fm a}). + refine (fix boundedSat (bound : nat) (fm : formula) {struct bound} + : option ({al : alist | satFormula fm (interp_alist al)} + + {forall a, ~satFormula fm a}) := + match bound with + | O => None + | S bound' => + if isNil fm + then Some (inleft _ (exist _ nil _)) + else match unitPropagate fm with + | Some (inleft (existT _ fm' (exist _ l _))) => + match boundedSat bound' fm' with + | None => None + | Some (inleft (exist _ al _)) => Some (inleft _ (exist _ (l :: al) _)) + | Some (inright _) => Some (inright _ _) + end + | Some (inright _) => Some (inright _ _) + | None => + let l := chooseSplit fm in + match setFormula fm l with + | inleft (exist _ fm' _) => + match boundedSat bound' fm' with + | None => None + | Some (inleft (exist _ al _)) => Some (inleft _ (exist _ (l :: al) _)) + | Some (inright _) => + match setFormula fm (negate l) with + | inleft (exist _ fm' _) => + match boundedSat bound' fm' with + | None => None + | Some (inleft (exist _ al _)) => Some (inleft _ (exist _ (negate l :: al) _)) + | Some (inright _) => Some (inright _ _) + end + | inright _ => Some (inright _ _) + end + end + | inright _ => + match setFormula fm (negate l) with + | inleft (exist _ fm' _) => + match boundedSat bound' fm' with + | None => None + | Some (inleft (exist _ al _)) => Some (inleft _ (exist _ (negate l :: al) _)) + | Some (inright _) => Some (inright _ _) + end + | inright _ => Some (inright _ _) + end + end + end + end); simpl; intros; subst; intuition; try generalize dependent (interp_alist al). + firstorder. + firstorder. + firstorder. + firstorder. + assert (satFormula fm (upd a0 l)); firstorder. + assert (satFormula fm (upd a0 l)); firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + firstorder. + destruct (satLit_dec l a); + [assert (satFormula fm (upd a l)) | assert (satFormula fm (upd a (negate l)))]; firstorder. + destruct (satLit_dec l a); + [assert (satFormula fm (upd a l)) | assert (satFormula fm (upd a (negate l)))]; firstorder. + destruct (satLit_dec l a); + [assert (satFormula fm (upd a l)) | assert (satFormula fm (upd a (negate l)))]; firstorder. + destruct (satLit_dec l a); + [assert (satFormula fm (upd a l)) | assert (satFormula fm (upd a (negate l)))]; firstorder. + destruct (satLit_dec l a); intuition eauto; + assert (satFormula fm (upd a l)); firstorder. + destruct (satLit_dec l a); intuition eauto; + assert (satFormula fm (upd a l)); firstorder. + firstorder. + firstorder. + destruct (satLit_dec l a); intuition eauto; + assert (satFormula fm (upd a (negate l))); firstorder. + destruct (satLit_dec l a); intuition eauto; + assert (satFormula fm (upd a (negate l))); firstorder. + destruct (satLit_dec l a); + [assert (satFormula fm (upd a l)) | assert (satFormula fm (upd a (negate l)))]; firstorder. +Admitted. + +Definition boundedSatSimple (bound : nat) (fm : formula) := + match boundedSat bound fm with + | None => None + | Some (inleft (exist _ a _)) => Some (Some a) + | Some (inright _) => Some None + end. + +(*Eval compute in boundedSatSimple 100 nil. +Eval compute in boundedSatSimple 100 (((true, 1%nat) :: nil) :: ((false, 1%nat) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((true, 0) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((true, 0) :: nil) :: ((false, 0) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((true, 0) :: (false, 1) :: nil) :: ((true, 1) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((true, 0) :: (false, 1) :: nil) :: ((true, 1) :: (false, 0) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((true, 0) :: (false, 1) :: nil) :: ((false, 0) :: (false, 0) :: nil) :: ((true, 1) :: nil) :: nil). +Eval compute in boundedSatSimple 100 (((false, 0) :: (true, 1) :: nil) :: ((false, 1) :: (true, 0) :: nil) :: nil).*) diff --git a/src/Smtpredicate.v b/src/Smtpredicate.v new file mode 100644 index 0000000..3594d53 --- /dev/null +++ b/src/Smtpredicate.v @@ -0,0 +1,5138 @@ +Require Import Coq.Logic.Decidable. +Require Import Coq.Structures.Equalities. +Require Import Coq.Numbers.Cyclic.Int63.Int63. +Require Import Coq.Sorting.Sorted. +Require Import Coq.Classes.RelationClasses. + +Require Import TVSMT.Coqlib. +Require Import TVSMT.Maps. +Require Import TVSMT.Errors. +Require Import TVSMT.Sat. +Require Import TVSMT.Common. +Require Import TVSMT.Predicate. +Require Import TVSMT.Hashtree. +Import TVSMT.Predicate.PredicateNotation. + +Require SMTCoq.SMTCoq. +Import SMTCoq.SMT_terms. +Import SMTCoq.SMT_terms.Atom. +Import SMTCoq.SMT_terms.Form. +Import SMTCoq.Trace. + +#[local] Open Scope error_monad_scope. +#[local] Open Scope predicate. + +#[export] Hint Unfold is_true : smtpred. + +Local Notation predicate := (@pred positive). + +Parameter pred_verit_unsat : + PArray.array Atom.atom -> + PArray.array Form.form -> + PArray.array Int63.int -> + (PArray.array Form.form + * PArray.array Atom.atom + * PArray.array Int63.int + * option (PArray.array Int63.int) + * Checker_Ext.certif) +. + +Notation latom := Int63.int (only parsing). +Notation _llit := Int63.int (only parsing). +Notation lfargs := (list _llit) (only parsing). + +Inductive lform : Type := +| LFatom (_:latom) +| LFtrue +| LFfalse +| LFnot2 (_:Int63.int) (_:_llit) +| LFand (_:lfargs) +| LFor (_:lfargs) +| LFimp (_:lfargs) +| LFxor (_:_llit) (_:_llit) +| LFiff (_:_llit) (_:_llit) +| LFite (_:_llit) (_:_llit) (_:_llit) +(* Bit-blasting predicate (theory of bit vectors): + bbT a [l1;...;ln] means [l1;...;ln] is the bitwise representation of a + (in little endian) + WARNING: this is a slight infringement of stratification + *) +| LFbbT (_:latom) (_:list _llit) +(* TODO: replace [list _lit] with [fargs] *) +. + +Fixpoint of_list' {A} (i: Z) (a: PArray.array A) (l: list A) := + match l with + | nil => a + | el :: b => + of_list' (i+1) (PArray.set a (Int63.of_Z i) el) b + end. + +Definition of_list {A} (d: A) (l: list A) := + of_list' 0 (PArray.make (Int63.of_Z (Zlength l)) d) l. + +Definition lform_to_form (f: lform): form := + match f with + | LFatom a => Fatom a + | LFtrue => Ftrue + | LFfalse => Ffalse + | LFnot2 i l => Fnot2 i l + | LFand args => Fand (of_list (Int63.of_Z 0) args) + | LFor args => For (of_list (Int63.of_Z 0) args) + | LFimp args => Fimp (of_list (Int63.of_Z 0) args) + | LFxor l1 l2 => Fxor l1 l2 + | LFiff l1 l2 => Fiff l1 l2 + | LFite l1 l2 l3 => Fite l1 l2 l3 + | LFbbT a b => FbbT a b + end. + +Definition form_to_lform (f: form): lform := + match f with + | Fatom a => LFatom a + | Ftrue => LFtrue + | Ffalse => LFfalse + | Fnot2 i l => LFnot2 i l + | Fand args => LFand (Misc.to_list args) + | For args => LFor (Misc.to_list args) + | Fimp args => LFimp (Misc.to_list args) + | Fxor l1 l2 => LFxor l1 l2 + | Fiff l1 l2 => LFiff l1 l2 + | Fite l1 l2 l3 => LFite l1 l2 l3 + | FbbT a b => LFbbT a b + end. + +Definition int_dec: forall a b: Int63.int, {a = b} + {a <> b}. +Proof. + intros. + destruct (Int63.eqb a b) eqn:?; [left|right]. + now apply Int63.eqb_correct. + unfold not; intros. apply Int63.eqb_complete in H. + now rewrite H in Heqb0. +Defined. + +Definition form_eq_dec : + forall d1 d2: lform, {d1 = d2} + {d1 <> d2}. +Proof. + intros. + pose proof int_dec. + pose proof (list_eq_dec H). + decide equality. +Defined. + +Definition atom_eq_dec : + forall d1 d2: atom, {d1 = d2} + {d1 <> d2}. +Proof. + intros. destruct (eqb d1 d2) eqn:?. apply eqb_spec in Heqb. auto. + right. unfold not; intros. apply Typ.not_is_true_eq_false in Heqb. + apply Heqb. unfold is_true. apply eqb_spec. trivial. +Defined. + +Module AtomDec <: MiniDecidableType. + Definition t := atom. + Definition eq_dec := atom_eq_dec. +End AtomDec. +Module AtomHash := Make_UDT(AtomDec). +Module Import AH := HashTree(AtomHash). +Module Import AH_Prop := HashTreeProperties(AtomHash). + +Module FormDec <: MiniDecidableType. + Definition t := lform. + Definition eq_dec := form_eq_dec. +End FormDec. +Module FormHash := Make_UDT(FormDec). +Module Import FH := HashTree(FormHash). +Module Import FH_Prop := HashTreeProperties(FormHash). + +Definition pred_list: Type := (predicate * list (positive * predicate)). + +Definition to_equiv (p: list (positive * predicate)) := + map (fun x => equiv (Pbase (fst x)) (snd x)) p. + +Definition eval_pred_list (p: pred_list) (a: PTree.t bool): option bool := + eval_hash_pred + (Pand (fst p) + (fold_left Pand (to_equiv (snd p)) Ptrue)) a. + +Fixpoint unnest_predicate (fresh: positive) (p: predicate) + : (pred_list * positive) := + match p with + | Ptrue => (Ptrue, nil, fresh) + | Pfalse => (Pfalse, nil, fresh) + | Pundef => (Pundef, nil, fresh) + | Pbase p' => (Pbase p', nil, fresh) + | Pand p1 p2 => + let '(p1', p1l, fresh1) := unnest_predicate fresh p1 in + let '(p2', p2l, fresh2) := unnest_predicate fresh1 p2 in + let nvar := Pbase fresh2 in + (nvar, (fresh2, (Pand p1' p2')) :: p2l ++ p1l, Pos.succ fresh2) + | Por p1 p2 => + let '(p1', p1l, fresh1) := unnest_predicate fresh p1 in + let '(p2', p2l, fresh2) := unnest_predicate fresh1 p2 in + let nvar := Pbase fresh2 in + (nvar, (fresh2, (Por p1' p2')) :: p2l ++ p1l, Pos.succ fresh2) + | Pimp p1 p2 => + let '(p1', p1l, fresh1) := unnest_predicate fresh p1 in + let '(p2', p2l, fresh2) := unnest_predicate fresh1 p2 in + let nvar := Pbase fresh2 in + (nvar, (fresh2, (Pimp p1' p2')) :: p2l ++ p1l, Pos.succ fresh2) + | Pnot p1 => + let '(p1', p1l, fresh1) := unnest_predicate fresh p1 in + let nvar := Pbase fresh1 in + (nvar, (fresh1, Pnot p1') :: p1l, Pos.succ fresh1) + end. + +Notation form_t := (PArray.array SMT_terms.Form.form). +Notation atom_t := (PArray.array SMT_terms.Atom.atom). +Notation root_t := (PArray.array Int63.int) (only parsing). + +Fixpoint all_preds (tree: PTree.t unit) (p: predicate): PTree.t unit := + match p with + | Ptrue | Pfalse | Pundef => tree + | Pbase p' => PTree.set p' tt tree + | Pand p1 p2 | Por p1 p2 | Pimp p1 p2 => all_preds (all_preds tree p1) p2 + | Pnot p1 => all_preds tree p1 + end. + +Definition of_P (a: positive) := Int63.of_Z (Zpos a - 2). +Definition to_P (a: Int63.int) := Z.to_pos (Int63.to_Z a + 2). + +Definition P_of_P (a: positive) := Int63.of_Z ((Zpos a - 2) * 2). +Definition N_of_P (a: positive) := Int63.of_Z ((Zpos a - 2) * 2 + 1). + +Definition of_map {A} (l: PMap.t A): PArray.array A := + PTree.fold (fun st i a => PArray.set st (of_P i) a) + (snd l) + (PArray.make (Int63.of_pos (max_key (snd l))) (fst l)). + +Definition of_tree {A} (d: A) (l: PTree.t A): option (PArray.array A) := + let max := ((max_key l) + 1)%positive in + if Z.pos max PArray.set st (of_P i) a) + l (PArray.make (of_P max) d)) + else None. + +Lemma max_key_empty : + forall A (m: PTree.t A), + PTree.bempty m = true -> + max_key m = 1%positive. +Proof. + intros. + rewrite PTree.bempty_correct in H. + unfold max_key. + rewrite PTree.elements_extensional with (n:= PTree.Leaf); auto. + intros. rewrite PTree.gleaf; auto. +Qed. + +Lemma max_key_list_in : forall l, + ((fold_right Pos.max 1%positive l) = 1%positive /\ l = nil) + \/ (In (fold_right Pos.max 1%positive l) l /\ l <> nil). +Proof. + induction l ; intros. + - auto. + - inv IHl. + + inv H. right. simpl. + split; [lia | congruence]. + + inv H. right. simpl. + split; [| congruence]. + edestruct Pos.max_spec with (n:= a) (m:= (fold_right Pos.max 1%positive l)). + * inv H. rewrite H3. right. auto. + * inv H. rewrite H3. left. auto. +Qed. + +Lemma max_key_list_not_nil : forall A l (m: PTree.t A), + (l <> nil) -> + (forall k v, m ! k = Some v -> In k l) -> + (forall k, In k l -> exists v, m ! k = Some v) -> + exists x, m ! (fold_right Pos.max 1%positive l) = Some x. +Proof. + intros A l. + generalize (max_key_list_in l). + intros [HCASE | HCASE]; intros. + - inv HCASE. congruence. + - inv HCASE. + eapply H1; eauto. +Qed. + +Lemma PTree_bempty_false_exists : forall A (m: PTree.t A), + PTree.bempty m = false -> + exists x v, m ! x = Some v. +Proof. + induction m ; intros. + - simpl in *. congruence. + - simpl in *. + destruct o. + + exists xH. simpl. eauto. + + apply andb_false_iff in H. + inv H. + * exploit IHm1; eauto. intros [x [v HSOME]]. + exists (xO x). eauto. + * exploit IHm2; eauto. intros [x [v HSOME]]. + exists (xI x). eauto. +Qed. + +Lemma max_key_in : + forall A (m: PTree.t A), + PTree.bempty m = false -> + exists x, m ! (max_key m) = Some x. +Proof. + unfold max_key. + intros * BEMPTY. + eapply max_key_list_not_nil with (l := (map fst (PTree.elements m))). + - intro. apply map_eq_nil in H. + generalize (PTree_bempty_false_exists _ m BEMPTY); eauto. + intros [x [v HGET]]. + eapply PTree.elements_correct in HGET. + rewrite H in *. inv HGET. + - intros. + eapply PTree.elements_correct in H. + apply in_map_iff. exists (k,v). auto. + - intros. + apply in_map_iff in H. destruct H as [[kk v] [Hfst Helt]]. subst. + simpl. eapply PTree.elements_complete in Helt. + eauto. +Qed. + +Lemma max_key_ext1 : + forall A (m: PTree.t A) m', + (forall x, m ! x = m' ! x) -> + PTree.bempty m = false -> + (max_key m <= max_key m')%positive. +Proof. + intros * ALL BEMPTY. + pose proof (max_key_in _ m BEMPTY) as X. inversion X as [x IN]; clear X. + pose proof IN as Y. rewrite ALL in Y. + assert (BEMPTY2: PTree.bempty m' = false). + { apply not_true_iff_false; unfold not; intros BEMPTY2. + apply not_true_iff_false in BEMPTY. apply BEMPTY. + apply PTree.bempty_correct; intros. + rewrite ALL. eapply PTree.bempty_correct in BEMPTY2; eauto. + } + pose proof (max_key_in _ m' BEMPTY2) as XX. inversion XX as [y IN2]; clear XX. + eapply max_key_correct; eauto. +Qed. + +Lemma max_key_ext2 : + forall A (m: PTree.t A) m', + (forall x, m ! x = m' ! x) -> + PTree.bempty m' = false -> + (max_key m' <= max_key m)%positive. +Proof. + intros * ALL BEMPTY. + pose proof (max_key_in _ m' BEMPTY) as X. inversion X as [x IN]; clear X. + pose proof IN as Y. rewrite <- ALL in Y. + assert (BEMPTY2: PTree.bempty m = false). + { apply not_true_iff_false; unfold not; intros BEMPTY2. + apply not_true_iff_false in BEMPTY. apply BEMPTY. + apply PTree.bempty_correct; intros. + rewrite <- ALL. eapply PTree.bempty_correct in BEMPTY2; eauto. + } + pose proof (max_key_in _ m BEMPTY2) as XX. inversion XX as [y IN2]; clear XX. + eapply max_key_correct; eauto. +Qed. + +Lemma max_key_ext : + forall A (m: PTree.t A) m', + (forall x, m ! x = m' ! x) -> + max_key m' = max_key m. +Proof. + unfold max_key; intros; erewrite PTree.elements_extensional; eauto. +Qed. + +Lemma max_key_map_le : + forall A B (m: PTree.t A) (m': PTree.t B), + (forall x, m ! x <> None <-> m' ! x <> None) -> + PTree.bempty m = false -> + (max_key m <= max_key m')%positive. +Proof. + intros * ALL BEMPTY. + assert (BEMPTY2: PTree.bempty m' = false). + { apply not_true_iff_false; unfold not; intros BEMPTY2. + apply not_true_iff_false in BEMPTY. apply BEMPTY. + apply PTree.bempty_correct; intros. + case_eq (m ! x); intros; auto. + specialize (ALL x). rewrite H in *. + rewrite PTree.bempty_correct in BEMPTY2. + specialize (BEMPTY2 x). + intuition congruence. + } + pose proof (max_key_in _ m BEMPTY) as X. inversion X as [x INx]; clear X. + assert (Y : exists y, m' ! (max_key m) = Some y). + { case_eq (m' ! (max_key m)); intros; eauto. + specialize (ALL (max_key m)). intuition congruence. + } + inversion Y as [y INy]. + try solve [eapply max_key_correct; eauto]. +Qed. + +Lemma ptree_map_same_keys: forall A B (f: positive -> A -> B) m, + forall x, m ! x <> None <-> (PTree.map f m) ! x <> None. +Proof. + split; intros. + - intro. elim H. + rewrite PTree.gmap in H0. + case_eq (m ! x); intros; auto. + rewrite H1 in H0; auto. + simpl in H0. congruence. + - intro. elim H. + rewrite PTree.gmap. + rewrite H0. auto. +Qed. + +Lemma max_key_map : + forall A B (m: PTree.t A) (f: positive -> A -> B), + max_key m = max_key (PTree.map f m). +Proof. + intros *. destruct (PTree.bempty m) eqn:Heqb. + - assert (BEMPTY2: PTree.bempty (PTree.map f m) = true). + { apply PTree.bempty_correct; intros. + eapply PTree.bempty_correct in Heqb; eauto. + rewrite PTree.gmap; eauto. + rewrite Heqb. auto. + } + rewrite ! max_key_empty; auto. + - assert (BEMPTY2: PTree.bempty (PTree.map f m) = false). + { apply not_true_iff_false; intro BEMPTY2. + apply not_true_iff_false in Heqb. apply Heqb. + apply PTree.bempty_correct; intros. + eapply PTree.bempty_correct with (x:= x) in BEMPTY2; eauto. + rewrite PTree.gmap in BEMPTY2; eauto. + unfold option_map in *. destruct_match; easy. + } + eapply Pos.le_antisym; eauto. + + eapply max_key_map_le; eauto. + eapply ptree_map_same_keys; eauto. + + eapply max_key_map_le; eauto. + split. + * intros. eapply ptree_map_same_keys; eauto. + * intros. eapply ptree_map_same_keys in H; eauto. +Qed. + +Lemma gt_1_map_iff : + forall A B (m: PTree.t A) (f: positive -> A -> B), + gt_1 m <-> gt_1 (PTree.map f m). +Proof. + unfold gt_1 in *; simplify. + split. + - simplify. + rewrite PTree.gmap in *. unfold option_map in *. + destruct (m ! x) eqn:?. eapply H. eauto. easy. + - simplify. + eapply H; eauto. + rewrite PTree.gmap. unfold option_map in *. + rewrite H0. easy. +Qed. + +(* From Coq.micromega Require Import RingMicromega QMicromega EnvRing Tauto ZifyComparison ZifyBool ZifyInt63. *) +(* Ltac Zify.zify_post_hook ::= Z.div_mod_to_equations. *) + +Lemma bounded_lt : + forall a b n, a < b -> 0 <= a < n -> 0 <= b < n -> a mod n < b mod n. +Proof. + intros; rewrite !Z.mod_small; lia. Qed. + +Lemma max_key_set_le : forall A (m: PTree.t A) k v, + (max_key m <= max_key (PTree.set k v m))%positive. +Proof. + intros. + case_eq (Pos.leb (max_key m) (max_key (PTree.set k v m)))%positive; intros. + - apply Pos.leb_le; auto. + - apply Pos.leb_gt in H; auto. + apply Pos.lt_gt in H. + exploit max_not_present; eauto. + intros. + assert (k <= max_key (PTree.set k v m))%positive. + { eapply max_key_correct; eauto. + rewrite PTree.gss; auto. + } + case_eq (PTree.bempty m); intros. + + exploit max_key_empty; eauto. + intros. rewrite H3 in *. + lia. + + exploit (max_key_in _ m); eauto. + intros [vmk HSome]. + rewrite PTree.gso in H0. + * congruence. + * intro; subst. + rewrite PTree.gss in H0. + congruence. +Qed. + +Lemma PArray_length_fold_set : forall A B a f (t: PTree.t A) (g: A -> B), + PArray.length (PTree.fold (fun st i a => PArray.set st (f i) (g a)) t a) + = PArray.length a. +Proof. + intros. + eapply PTree_Properties.fold_rec; eauto. + intros. + rewrite PArray.length_set. + auto. +Qed. + +Lemma of_pos_rec_inj: forall p1 p2 + (PVALID: 0 <= Z.pos p1 < wB) + (KVALID: 0 <= Z.pos p2 < wB), + (φ (of_pos_rec size p1) = φ (of_pos_rec size p2))%int63 -> + p1 = p2. +Proof. + unfold wB. intros. + rewrite ! of_pos_rec_spec in H; auto. + rewrite ! Z.mod_small in H; try lia. +Qed. + +Lemma of_P_inj : forall k p + (PVALID: 0 <= Z.pos p - 2 < wB) + (KVALID: 0 <= Z.pos k - 2 < wB), + of_P k = of_P p -> k = p. +Proof. + unfold of_P, of_Z, wB, of_pos; intros. + do 2 destruct_match; try lia. + - generalize (of_pos_rec_spec size (Nat.le_refl _) p0). + rewrite Z.mod_small; auto. intros. + rewrite <- H in H0. + rewrite to_Z_0 in H0. + lia. + - generalize (of_pos_rec_spec size (Nat.le_refl _) p0). + rewrite Z.mod_small; auto. intros. + rewrite H in H0. + rewrite to_Z_0 in H0. + lia. + - generalize (of_pos_rec_spec size (Nat.le_refl _) p0). + generalize (of_pos_rec_spec size (Nat.le_refl _) p1). + rewrite !Z.mod_small; auto. intros. + assert (p0 = p1). + { eapply of_pos_rec_inj; eauto. + congruence. + } + lia. +Qed. + +Definition tree_a_wf (t: PTree.t atom) := + 1 < Zpos (max_key t) < wB + /\ gt_1 t + /\ t ! 2 = Some (Acop CO_xH) + /\ t ! 3 = Some (Auop UO_Zpos (of_P 2)) + /\ t ! 4 = Some (Auop UO_Zopp (of_P 3)) +. + +Definition tree_f_wf (t: PTree.t lform) := + 1 < Zpos (max_key t) + /\ (Zpos (max_key t) - 2) * 2 + 1 < wB + /\ gt_1 t + /\ t ! 2 = Some (LFtrue) + /\ t ! 3 = Some (LFfalse) +. + +(* Abstracting over the type of elements of the tree *) +Definition tree_key_wf A (t: PTree.t A) := + 1 < Zpos (max_key t) + /\ ((Zpos (max_key t) - 2) * 2 + 1 < wB + \/ Zpos (max_key t) < wB) + /\ gt_1 t. + +Lemma of_tree_of_P : + forall A d t p (v: A) a + (TAWF: tree_key_wf _ t), + gt_1 t -> + t ! p = Some v -> + of_tree d t = Some a -> + PArray.get a (of_P p) = v. +Proof. + unfold of_tree. + intros. destruct_match; try discriminate. inv H1. + assert + (X: 0 <= Z.pos (max_key t) - 1 <= + to_Z (PArray.length (PTree.fold (fun st i a => PArray.set st (of_P i) a) t + (PArray.make (of_P (max_key t + 1)) d)))). + { erewrite PArray_length_fold_set; eauto. + rewrite PArray.length_make. + unfold PArray.max_length. + unfold wB, size in Heqb. + inv TAWF. + destruct_match; try lia. + - unfold of_P. rewrite of_Z_spec. + rewrite Z.mod_small; try lia. + - rewrite Misc.max_int_wB. + unfold wB in *. lia. + } + unfold gt_1 in *. + generalize dependent H. + generalize dependent Heqb. + generalize dependent v. + generalize dependent p. + generalize dependent X. + apply PTree_Properties.fold_rec; eauto. intros. + - rewrite <- H in H1. + eapply H0; auto. + + erewrite max_key_ext; eauto. + + erewrite max_key_ext; eauto. + + intros. rewrite H in H3; eauto. + - intros. + now rewrite PTree.gempty in H0. + - intros. + destruct (peq k p); subst. + + rewrite PTree.gss in H2. inv H2. + rewrite PArray.get_set_same; auto. + apply Int63.ltb_spec. + rewrite PArray.length_set in X. + destruct X as [XX X]. + eapply Z.lt_le_trans; eauto. + * unfold of_P. rewrite of_Z_spec. + assert (p <= max_key (PTree.set p v0 m))%positive. + { eapply max_key_correct; eauto. + rewrite PTree.gss; auto. + } + rewrite Z.mod_small; try lia. + inv TAWF. inv H5. unfold gt_1 in *. + specialize (H7 _ _ H0). + apply max_key_correct with (hi:= p) in H0; eauto. + apply Pos2Z.pos_le_pos in H0. lia. + + rewrite PTree.gso in H2; auto. + rewrite PArray.get_set_other. + * apply H1; auto. + -- rewrite PArray.length_set in X. + destruct X as [XX X]. + split; [lia |]. + eapply Z.le_trans; eauto. + assert (max_key m <= max_key (PTree.set k v m))%positive + by (eapply max_key_set_le; eauto). + lia. + -- rewrite PArray.length_set in X. + assert (max_key m <= max_key (PTree.set k v m))%positive + by (eapply max_key_set_le; eauto). + rewrite Z.ltb_lt in *. + lia. + -- intros. eapply H3; eauto. + erewrite PTree.gso; eauto. + congruence. + * intro; subst. + eapply of_P_inj in H4; eauto. + apply Z.ltb_lt in Heqb. + rewrite PArray.length_set in X. + -- assert (1 < p)%positive. + { eapply H3; eauto. + rewrite PTree.gso; eauto. + } + assert (p <= max_key (PTree.set k v m))%positive. + { eapply max_key_correct; eauto. + rewrite PTree.gso; eauto. + } + lia. + -- assert (1 < k)%positive. + { eapply H3; eauto. + rewrite PTree.gss; auto. + } + assert (k <= max_key (PTree.set k v m))%positive. + { eapply max_key_correct; eauto. + rewrite PTree.gss; auto. + } + apply Z.ltb_lt in Heqb; auto. + rewrite PArray.length_set in X. + lia. +Qed. + +Definition to_smt_i ah p := + match p with + | Pbase p => match AH.find_tree (Aapp (of_P p) nil) ah with + | Some pi => + if (1 None + end + | Ptrue => Some (Int63.of_Z 1) + | Pfalse => Some (Int63.of_Z 2) + | _ => None + end. + +Fixpoint get_index_acc {A} (dec: A -> A -> bool) (l : list A) + (a : A) (acc : nat) {struct l} : option nat := + match l with + | nil => None + | x :: m => if dec x a then Some acc else get_index_acc dec m a (acc + 1)%nat + end. + +Definition get_index {A} dec (l : list A) (a : A) := get_index_acc dec l a 0. + +Definition form_eqb (f1 f2: form) := + match f1, f2 with + | Fatom i1, Fatom i2 => Int63.eqb i1 i2 + | Ftrue, Ftrue => true + | Ffalse, Ffalse => true + | Fnot2 a1 b1, Fnot2 a2 b2 + | Fiff a1 b1, Fiff a2 b2 + | Fxor a1 b1, Fxor a2 b2 => Int63.eqb a1 a2 && Int63.eqb b1 b2 + | Fand args1, Fand args2 + | For args1, For args2 + | Fimp args1, Fimp args2 => list_beq Int63.eqb (Misc.to_list args1) (Misc.to_list args2) + | Fite a1 b1 c1, Fite a2 b2 c2 => Int63.eqb a1 a2 && Int63.eqb b1 b2 && Int63.eqb c1 c2 + | _, _ => false + end. + +Definition get_fi (lf: list form) (fresh: Z) (f: form): Z * Z * list form := + match get_index form_eqb lf f with Some i => (Z.of_nat i, fresh, lf) | _ => (fresh, Z.succ fresh, lf ++ f::nil) end. + +Definition get_ai (lf: list atom) (fresh: Z) (f: atom): Z * Z * list atom := + match get_index Atom.eqb lf f with Some i => (Z.of_nat i, fresh, lf) | _ => (fresh, Z.succ fresh, lf ++ f::nil) end. + +Definition to_smt_and hf ha p p1 p2 := + let (a1, ha1) := AH.hash_value 1%positive (Abop BO_Zlt p1 p2) ha in + let (a2, ha2) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p p1) ha1 in + let (a3, ha3) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p p2) ha2 in + let (i1, hf1) := FH.hash_value 1%positive (LFatom (of_P a1)) hf in + let (i2, hf2) := FH.hash_value 1%positive (LFatom (of_P a2)) hf1 in + let (i3, hf3) := FH.hash_value 1%positive (LFatom (of_P a3)) hf2 in + let (i4, hf4) := FH.hash_value 1%positive (LFite (P_of_P i1) (P_of_P i2) (P_of_P i3)) hf3 in + (hf4, ha3, i4). + +Definition to_smt_or hf ha p p1 p2 := + let '(a1, ha1) := AH.hash_value 1%positive (Abop BO_Zlt p1 p2) ha in + let '(a2, ha2) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p p2) ha1 in + let '(a3, ha3) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p p1) ha2 in + let '(i1, hf1) := FH.hash_value 1%positive (LFatom (of_P a1)) hf in + let '(i2, hf2) := FH.hash_value 1%positive (LFatom (of_P a2)) hf1 in + let '(i3, hf3) := FH.hash_value 1%positive (LFatom (of_P a3)) hf2 in + let '(i4, hf4) := FH.hash_value 1%positive (LFite (P_of_P i1) (P_of_P i2) (P_of_P i3)) hf3 in + (hf4, ha3, i4). + +Definition to_smt_imp hf ha p p1 p2 := + let '(a1, ha1) := AH.hash_value 1%positive (Abop BO_Zminus (of_P 3) p1) ha in + let '(a2, ha2) := AH.hash_value 1%positive (Abop BO_Zplus (of_P a1) p2) ha1 in + let '(a3, ha3) := AH.hash_value 1%positive (Abop BO_Zlt (of_P 3) (of_P a2)) ha2 in + let '(a4, ha4) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p (of_P 3)) ha3 in + let '(a5, ha5) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p (of_P a2)) ha4 in + let '(i1, hf1) := FH.hash_value 1%positive (LFatom (of_P a3)) hf in + let '(i2, hf2) := FH.hash_value 1%positive (LFatom (of_P a4)) hf1 in + let '(i3, hf3) := FH.hash_value 1%positive (LFatom (of_P a5)) hf2 in + let '(i4, hf4) := FH.hash_value 1%positive (LFite (P_of_P i1) (P_of_P i2) (P_of_P i3)) hf3 in + (hf4, ha5, i4). + +Definition to_smt_not hf ha p p1 := + let '(a1, ha1) := AH.hash_value 1%positive (Auop UO_Zopp p1) ha in + let '(a2, ha2) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p (of_P a1)) ha1 in + let '(i1, hf1) := FH.hash_value 1%positive (LFatom (of_P a2)) hf in + (hf1, ha2, i1). + +Definition to_smt_l (el: positive * predicate) (state: option (list Int63.int * FH.hash_tree * AH.hash_tree)) := + match state with + | Some (r, f, a) => + let to_smt_i := to_smt_i a in + match el with + | (p, Pand p1 p2) => + match to_smt_i (Pbase p), to_smt_i p1, to_smt_i p2 with + | Some p', Some p1', Some p2' => + let '(f', a', r') := to_smt_and f a p' p1' p2' in + Some (r++(P_of_P r'::nil), f', a') + | _, _, _ => None + end + | (p, Por p1 p2) => + match to_smt_i (Pbase p), to_smt_i p1, to_smt_i p2 with + | Some p', Some p1', Some p2' => + let '(f', a', r') := to_smt_or f a p' p1' p2' in + Some (r++(P_of_P r'::nil), f', a') + | _, _, _ => None + end + | (p, Pimp p1 p2) => + match to_smt_i (Pbase p), to_smt_i p1, to_smt_i p2 with + | Some p', Some p1', Some p2' => + let '(f', a', r') := to_smt_imp f a p' p1' p2' in + Some (r++(P_of_P r'::nil), f', a') + | _, _, _ => None + end + | (p, Pnot p1) => + match to_smt_i (Pbase p), to_smt_i p1 with + | Some p', Some p1' => + let '(f', a', r') := to_smt_not f a p' p1' in + Some (r++(P_of_P r'::nil), f', a') + | _, _ => None + end + | _ => None + end + | _ => None + end. + +Definition declare_atoms_with_bounds + (a: Z * list Int63.int * FH.hash_tree * AH.hash_tree) := + let '(num, cr, cf, ca) := a in + let (happ, ca_inter) := AH.hash_value 1%positive (Aapp (Int63.of_Z num) nil) ca in + let '((hlt1, hlt2), ca') := + AH.hash_value2 1%positive + (Abop BO_Zle (of_P 4) (of_P happ), + Abop BO_Zge (of_P 3) (of_P happ)) ca_inter in + let '((halt1, halt2), cf_inter) := + FH.hash_value2 1%positive (LFatom (of_P hlt1), LFatom (of_P hlt2)) cf in + let (hlfand, cf') := + FH.hash_value 1%positive (LFand (P_of_P halt1 :: P_of_P halt2 :: nil)) cf_inter in + (num + 1, cr ++ (P_of_P hlfand :: nil), cf', ca'). + +Definition check_tree_f_wf (f: PTree.t lform) := + let max := Z.pos (max_key f) in + (1 true | _ => false end + && match f ! 2 with Some LFtrue => true | _ => false end + && match f ! 3 with Some LFfalse => true | _ => false end. + +Definition check_tree_a_wf (a: PTree.t atom) := + let max := Z.pos (max_key a) in + (1 true | _ => false end + && match a ! 2 with Some (Acop CO_xH) => true | _ => false end + && match a ! 3 with Some (Auop UO_Zpos v) => (v =? 0)%int63 | _ => false end + && match a ! 4 with Some (Auop UO_Zopp v) => (v =? 1)%int63 | _ => false end. + +Definition to_smt (max: positive) (p: pred_list): option (root_t * form_t * atom_t) := + let (p', pl) := p in + let '(_, init_forms) := FH.hash_value2 1%positive (LFtrue, LFfalse) FH.empty in + let (hpone, ia_inter1) := AH.hash_value 1%positive (Acop CO_xH) AH.empty in + let (hzone, ia_inter2) := AH.hash_value 1%positive (Auop UO_Zpos (of_P hpone)) ia_inter1 in + let (hzmone, init_atoms) := AH.hash_value 1%positive (Auop UO_Zopp (of_P hzone)) ia_inter2 in + let '(v, r_defs, form_defs, atom_defs) := + fold_right (fun _ => declare_atoms_with_bounds) + (0, nil, init_forms, init_atoms) + (repeat tt (Pos.to_nat max - 1)) + in + match + fold_right to_smt_l (Some (r_defs, form_defs, atom_defs)) pl, + to_smt_i atom_defs p' + with + | Some (r, f, a), Some p'' => + let '(a1, fin_a) := AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) p'' (Int63.of_Z 1)) a in + let '(f1, fin_f) := FH.hash_value 1%positive (LFatom (of_P a1)) f in + let fin_r := r ++ (N_of_P f1)::nil in + match + of_tree Ftrue (PTree.map (fun _ => lform_to_form) fin_f), + of_tree (Acop CO_xH) fin_a + with + | Some fin_f', Some fin_a' => + if check_form fin_f' + && check_atom fin_a' + && check_tree_f_wf fin_f + && check_tree_a_wf fin_a + && (v None + end + | _, _ => None + end. + +Fixpoint check_list {A} (comp: A -> A -> bool) (a b: list A) := + match b, a with + | _, nil => true + | x::xs, y::ys => comp x y && check_list comp ys xs + | _, _ => false + end. + +Definition check_array {A} (comp: A -> A -> bool) a b := + check_list comp (Misc.to_list a) (Misc.to_list b). + +Definition single_form_not_bv f := + match f with + | Fatom _ | Ftrue | Ffalse | Fnot2 _ _ | Fand _ + | For _ | Fxor _ _ | Fiff _ _ | Fite _ _ _ => true + | FbbT _ _ | Fimp _ => false + end. + +Definition check_atoms_in_form (l: Int63.int) (f: form) := + match f with + | Fatom a => Int63.ltb a l + | _ => true + end. + +Definition check_local_form la f := + Form.check_form f + && Misc.aforallbi (fun i f' => single_form_not_bv f' + && check_atoms_in_form la f') f. + +Definition bool_typ_compdec := SMT_classes.Typ_compdec bool SMT_classes_instances.bool_compdec. +Definition Z_typ_compdec := SMT_classes.Typ_compdec Z SMT_classes_instances.Z_compdec. + +Definition t_i_empty := PArray.make (Int63.of_Z 0) Z_typ_compdec. + +Definition Ztrue := 1. +Definition Zfalse := -1. +Definition Zundef := 0. + +Definition Zand (a b: Z) := + if a 1 + | Some false => -1 + | None => 0 + end. + +Definition t_func_mk: Z -> tval t_i_empty := fun b => Tval t_i_empty (nil, Typ.TZ) b. + +(* Definition asgn_transl' (max: positive) (a: PTree.t bool) := *) +(* Misc.foldi (fun i na => PArray.set na i (t_func_mk (Zval (a!(to_P i))))) *) +(* (Int63.of_Z 0) (of_P max) *) +(* (PArray.make (of_P max) (t_func_mk 0)). *) + +Inductive asgn_transl_spec (a: PTree.t bool) (res: PArray.array (tval t_i_empty)) := +| asgn_transl_rule : + (forall i, 0 <= Z.pos i - 2 < wB -> PArray.get res (of_P i) = t_func_mk (Zval (a!i))) -> + PArray.default res = t_func_mk 0 /\ (forall i, exists n, PArray.get res i = t_func_mk n) -> + asgn_transl_spec a res. + +Definition asgn_transl (a: PTree.t bool) := + let max := max_key a in + PTree.fold (fun na i el => + PArray.set na (of_P i) (t_func_mk (Zval (Some el))) + ) a (PArray.make (of_P (max + 1)) (t_func_mk 0)). + +Lemma PArray_set_out_of_bound: forall A (a: PArray.array A) k d i, + (k + PArray.get (PArray.set a k d) i = PArray.get a i. +Proof. + intros. + unfold PArray.set in *. + destruct a. simpl PArray.length in *. + destruct_match. + - auto. + - rewrite Misc.ltb_negb_geb in H. + rewrite negb_false_iff in H. + congruence. +Qed. + +Lemma PArray_default_fold_set : forall A B (m: PTree.t A) f (g: A -> B) array, + PArray.default (PTree.fold (fun st i a => PArray.set st (f i) (g a)) m array) = + PArray.default array. +Proof. + intros. + eapply PTree_Properties.fold_rec; eauto. + intros. rewrite PArray.default_set; auto. +Qed. + +Lemma asgn_transl_OOB: forall a k, + (k + PArray.get (asgn_transl a) k = t_func_mk 0. +Proof. + unfold asgn_transl. + intros. + rewrite PArray.get_outofbound; auto. + rewrite PArray_default_fold_set; auto. +Qed. + +Lemma exists_asgn: + forall a, + forall (MAX: Z.pos (max_key a) - 2 < wB - 1), + (forall k x, a ! k = Some x -> (0 <= Z.pos k - 2 < wB)%int63) -> + exists res, asgn_transl_spec a res. +Proof. + intros. exists (asgn_transl a). + constructor. + intros i. + case_eq (of_P i to_Z (State.Lit.blit x) + match a' with + | Anop _ _ => false + | Aapp _ (_::_) => false + | _ => true + end) a. + +Fixpoint pred_in c p := + match p with + | Ptrue | Pfalse | Pundef => false + | Pbase c2 => (c =? c2)%positive + | Pand p1 p2 | Por p1 p2 | Pimp p1 p2 => pred_in c p1 || pred_in c p2 + | Pnot p' => pred_in c p' + end. + +Definition check_smt (p: predicate): bool := + let a_preds1 := + (max_hash_pred p + 1)%positive + in + let '(p', a_preds) := unnest_predicate a_preds1 p in + match to_smt a_preds p' with + | Some (d, forms, atoms) => + let '(nf, na, roots, used_roots, cert) := pred_verit_unsat atoms forms d in + check_array form_eqb forms nf + && check_array Atom.eqb atoms na + && check_local_form (PArray.length atoms) forms + && Atom.check_atom na + && Atom.wt _ empty_tfunc na + && check_wt_extra na + && (check_root_index forms d && negb (pred_in 1 p) && (Z.pos a_preds false + end. + +Definition mutate1_p (p: predicate): predicate := + match p with + | Pand p1 p2 => Pand p2 p1 + | Por p1 p2 => Por p2 p1 + | Pimp (Pand p1 p2) b => Pimp (Pand p2 p1) b + | Pimp (Por p1 p2) b => Pimp (Por p2 p1) b + | Pimp b (Pand p1 p2) => Pimp b (Pand p2 p1) + | Pimp b (Por p1 p2) => Pimp b (Por p2 p1) + | Pnot (Pand p1 p2) => Pnot (Pand p2 p1) + | Pnot (Por p1 p2) => Pnot (Por p2 p1) + | p => p + end. + +Fixpoint has_Pbase (p: predicate) := + match p with + | Ptrue | Pfalse | Pundef => false + | Pbase _ => true + | Pand a b | Por a b | Pimp a b => has_Pbase a || has_Pbase b + | Pnot a => has_Pbase a + end. + +Definition check_smt_total (p: predicate): bool := + if check_smt p then true + else if check_smt (mutate1_p p) then true + else false. + + (** We need a higher level specification for this function so that everything is easier to prove. *) + + Definition upd_asgn := + fold_left + (fun a' el => + match eval_hash_pred (snd el) a' with + | Some p_b => PTree.set (fst el) p_b a' + | _ => a' + end). + + Inductive upd_asgn_spec: list (positive * predicate) -> PTree.t bool -> PTree.t bool -> Prop := + | upd_asgn_intros x : + upd_asgn_spec nil x x + | upd_asgn_cons_Some x xs t1 t2 b: + eval_hash_pred (snd x) t1 = Some b -> + upd_asgn_spec xs (PTree.set (fst x) b t1) t2 -> + upd_asgn_spec (x :: xs) t1 t2 + | upd_asgn_cons_None x xs t1 t2: + eval_hash_pred (snd x) t1 = None -> + upd_asgn_spec xs t1 t2 -> + upd_asgn_spec (x :: xs) t1 t2. + + (* #[local] Open Scope positive. *) + (* Compute PTree.get 3 (upd_asgn ((3, Pbase 2)::(2, Pbase 1)::nil) (PTree.set 1 false (PTree.empty _))). *) + + (* Lemma upd_asgn_cons : *) + (* forall p form r a, *) + (* (forall pb, pred_In pb form -> ~ In pb (map fst r)) -> *) + (* (upd_asgn ((p, form)::r) a) ! p = eval_hash_pred form a. *) + (* Proof. *) + + (* Lemma upd_asgn_cons_form : *) + (* forall p form r a, *) + (* eval_hash_pred (Pbase p) (upd_asgn ((p, form)::r) a) = eval_hash_pred form a. *) + (* Proof. *) + + (* Lemma wt_correct' : *) + (* v_type Typ.type (Typ.interp t_i_empty) (PArray.get (Atom.t_interp t_i_empty a' atoms) i) *) + + Ltac ldestruct_match := + match goal with + | [ H: context[match ?x with | _ => _ end] |- _ ] => destruct x eqn:? + end. + + Lemma wt_correct' : + forall a' a (atoms: atom_t) + (ATOMSNONOP: forall i x l, PArray.get atoms i <> Anop x l) + (ATOMSAPP: forall i n l, PArray.get atoms i = Aapp n l -> l = nil) + (GETA1: forall i, exists n, PArray.get a' i = t_func_mk n) + (GETA2: forall i, exists n, PArray.get a i = t_func_mk n), + forall i, + v_type Typ.type (Typ.interp t_i_empty) + (PArray.get (Atom.t_interp t_i_empty a atoms) i) = + v_type Typ.type (Typ.interp t_i_empty) (PArray.get (Atom.t_interp t_i_empty a' atoms) i). + Proof. + intros. generalize dependent i. + pose proof (Atom.length_t_interp t_i_empty a atoms) as Y1. + pose proof (Atom.length_t_interp t_i_empty a' atoms) as Y2. + generalize dependent Y2. + generalize dependent Y1. + apply Misc.foldi_ind2 with + (P := (fun x b b' => + PArray.length b = PArray.length atoms -> + PArray.length b' = PArray.length atoms -> + (forall i, + v_type Typ.type (Typ.interp t_i_empty) (PArray.get b i) + = v_type Typ.type (Typ.interp t_i_empty) (PArray.get b' i)))). + + apply Misc.leb_0. + + intros. rewrite PArray.get_make in *. crush. + + intros. + rewrite PArray.length_set in *. + pose proof H0 as LT1. apply Int63.ltb_spec in LT1. + pose proof H as LE2. apply Int63.leb_spec in LE2. + destruct (int_dec i i0); subst. + rewrite !PArray.get_set_same by (apply Int63.ltb_spec; try (rewrite H3); try rewrite H2; lia); auto. + specialize (H1 H2 H3). assert (EQ1:=H2). assert (EQ2:=H3). clear H2. clear H3. + destruct (PArray.get atoms i0) eqn:ATOMS; subst; simplify. + - unfold interp_cop; repeat destruct_match; subst; auto. + - assert (HEQ := H1 i). unfold interp_uop; destruct_match; unfold apply_unop; + repeat destruct_match; auto; + setoid_rewrite Heqb0 in HEQ; simplify; + setoid_rewrite Heqb in HEQ; simplify; subst; rewrite Heqc in Heqc0; discriminate. + - assert (HEQ1 := H1 i). assert (HEQ2 := H1 i1). + unfold interp_bop. destruct_match; + unfold apply_binop; repeat destruct_match; auto; + try (setoid_rewrite Heqb4 in HEQ2; setoid_rewrite Heqb2 in HEQ2; + setoid_rewrite Heqb3 in HEQ1; setoid_rewrite Heqb1 in HEQ1; + simplify; subst; try rewrite Heqc0 in *; try rewrite Heqc1 in *; + try rewrite Heqc2 in *; try rewrite Heqc3 in *; discriminate). + - assert (HEQ1 := H1 i). assert (HEQ2 := H1 i1). assert (HEQ3 := H1 i2). + unfold interp_top. destruct_match; subst. unfold apply_terop. + repeat destruct_match; auto; + setoid_rewrite Heqb4 in HEQ3; setoid_rewrite Heqb2 in HEQ1; setoid_rewrite Heqb3 in HEQ2; + setoid_rewrite Heqb1 in HEQ3; setoid_rewrite Heqb in HEQ1; setoid_rewrite Heqb0 in HEQ2; + simplify; subst; try rewrite Heqc4 in *; try rewrite Heqc3 in *; try rewrite Heqc2 in *; + try rewrite Heqc1 in *; try rewrite Heqc0 in *; try rewrite Heqc in *; discriminate. + - exfalso. eapply ATOMSNONOP; eassumption. + - assert (HEQ1 := H1 i). specialize (GETA1 i). specialize (GETA2 i). + inv GETA1. inv GETA2. + repeat destruct_match. assert (l = nil) by (eapply ATOMSAPP; eassumption); subst. + cbn [map]. unfold t_func_mk in *. unfold Tval in *. inv H2. simpl. + inv H3. auto. + - rewrite !PArray.get_set_other + by (auto; apply Int63.ltb_spec; try (rewrite H3); try rewrite H2; lia); + eauto. + Qed. + + Lemma wt_correct2 : + forall atoms a a' + (ATOMSNONOP: forall i x l, PArray.get atoms i <> Anop x l) + (ATOMSAPP: forall i n l, PArray.get atoms i = Aapp n l -> l = nil) + (GETA1: forall i, exists n, PArray.get a' i = t_func_mk n) + (GETA2: forall i, exists n, PArray.get a i = t_func_mk n), + is_true (wt t_i_empty a atoms) -> + is_true (wt t_i_empty a' atoms). + Proof. + unfold is_true, wt; intros. + apply Misc.aforallbi_spec; intros i LTB. + eapply Misc.aforallbi_spec in H; [|eassumption]. + unfold check_aux in *. + destruct (PArray.get atoms i) eqn:?; unfold get_type' in *. + - erewrite wt_correct'; eassumption. + - destruct_match. simplify. + setoid_rewrite wt_correct' with (a':=a); eauto. + rewrite H0. rewrite H1. auto. + - repeat destruct_match; simplify. + setoid_rewrite wt_correct' with (a':=a); eauto. + rewrite H1. rewrite H. rewrite H2. auto. + - repeat destruct_match; simplify. + setoid_rewrite wt_correct' with (a':=a); eauto. + rewrite H1. rewrite H0. rewrite H2. rewrite H3. auto. + - exfalso; eapply ATOMSNONOP; eauto. + - repeat destruct_match; crush. eapply ATOMSAPP in Heqa0; subst. simplify. + assert (GETB1:=GETA1). assert (GETB2:=GETA2). + specialize (GETA1 i0). specialize (GETA2 i0). inv GETA1. inv GETA2. + rewrite H in Heqf. rewrite H2 in Heqf0. unfold t_func_mk, Tval in *. + simplify. erewrite wt_correct'; eauto. + Qed. + + Lemma wf_empty_tfunc : + forall i, + exists n, PArray.get empty_tfunc i = t_func_mk n. + Proof. unfold empty_tfunc. intros. eexists. rewrite PArray.get_make. eauto. Qed. + + Lemma to_P_of_P : + forall p, 1 < Z.pos p < wB -> to_P (of_P p) = p. + Proof. + unfold to_P, of_P; intros * LT. + rewrite of_Z_spec. rewrite !Z.mod_small by lia. + assert (forall x, x - 2 + 2 = x) by lia. + rewrite H; auto. + Qed. + +Lemma not_int63_lt_0 : forall i, 0 <= Int63.to_Z i. +Proof. + intro i. + assert (~ Int63.to_Z i < 0). + { unfold not; intros. rewrite <- Int63.to_Z_0 in H. + apply Int63.ltb_spec in H. now apply Misc.ltb_0 in H. } + lia. +Qed. + + Lemma of_P_to_P : + forall p, of_P (to_P p) = p. + Proof. + unfold to_P, of_P; intros. rewrite Z2Pos.id. + assert (forall x, x + 2 - 2 = x) by lia. rewrite H. + now rewrite of_to_Z. + pose proof (not_int63_lt_0 p). lia. + Qed. + + Lemma wf_asgn_spec : + forall a a', + asgn_transl_spec a a' -> + forall i, exists n, PArray.get a' i = t_func_mk n. + Proof. inversion 1; intros. inv H1. eauto. Qed. + + Lemma wt_correct : + forall atoms a a' + (ATOMSNONOP: forall i x l, PArray.get atoms i <> Anop x l) + (ATOMSAPP: forall i n l, PArray.get atoms i = Aapp n l -> l = nil), + is_true (wt _ empty_tfunc atoms) -> + asgn_transl_spec a a' -> + is_true (wt _ a' atoms). + Proof. + intros. + pose proof (wf_asgn_spec _ _ H0). + pose proof wf_empty_tfunc. + eapply wt_correct2; eauto. + Qed. + +Section SMT_PROOF. + + Fixpoint Zeval (a: PTree.t bool) (p: predicate) := + match p with + | Ptrue => 1 + | Pfalse => -1 + | Pundef => 0 + | Pbase p' => Zval (a ! p') + | Pnot p' => Znot (Zeval a p') + | Pand p1 p2 => Zand (Zeval a p1) (Zeval a p2) + | Por p1 p2 => Zor (Zeval a p1) (Zeval a p2) + | Pimp p1 p2 => Zimp (Zeval a p1) (Zeval a p2) + end. + + Lemma Zeval_correct : + forall a p, + Zeval a p = Zval (eval_hash_pred p a). + Proof. + induction p; crush; + repeat (destruct_match; simplify); + try (rewrite IHp1; rewrite IHp2); + auto; rewrite IHp; auto; discriminate. + Qed. + +End SMT_PROOF. + +Lemma check_local_form_spec : + forall l f, + is_true (check_local_form l f) -> + (forall i a b, PArray.get f i <> FbbT a b) + /\ (forall i a, PArray.get f i <> Fimp a) + /\ (forall i a, PArray.get f i = Fatom a -> is_true (Int63.ltb a l)) + /\ is_true (check_form f). +Proof. + unfold check_local_form, is_true; simplify; auto. + - unfold not; intros. + destruct (Int63.ltb i (PArray.length f)) eqn:Heqb. + + eapply Misc.aforallbi_spec in H1; try eassumption. now rewrite H in H1. + + unfold check_form in H0. simplify. unfold is_Ffalse, is_Ftrue in *. + repeat (destruct_match; try discriminate). + rewrite PArray.get_outofbound in H; auto. now rewrite Heqf1 in H. + - unfold not; intros. + destruct (Int63.ltb i (PArray.length f)) eqn:Heqb. + + eapply Misc.aforallbi_spec in H1; try eassumption. now rewrite H in H1. + + unfold check_form in H0. simplify. unfold is_Ffalse, is_Ftrue in *. + repeat (destruct_match; try discriminate). + rewrite PArray.get_outofbound in H; auto. now rewrite Heqf1 in H. + - intros. + destruct (Int63.ltb i (PArray.length f)) eqn:Heqb. + + eapply Misc.aforallbi_spec in H1; try eassumption. simplify. + unfold check_atoms_in_form in H3. now rewrite H in H3. + + unfold check_form in H0. simplify. unfold is_Ffalse, is_Ftrue in *. + repeat (destruct_match; try discriminate). + rewrite PArray.get_outofbound in H by auto. now rewrite Heqf1 in H. +Qed. + +Lemma to_list_corr' : + forall A a, + (fun f b => + (forall i, (Int63.to_Z i < Int63.to_Z f) -> + PArray.get a i = nth (Z.to_nat (Int63.to_Z i)) (rev b) (PArray.default a)) + /\ Int63.to_Z f = Zlength b) + (PArray.length a) + (Misc.foldi (fun i (l : list A) => PArray.get a i :: l) + (Int63.of_Z 0) + (PArray.length a) nil). +Proof. + intros. apply Misc.foldi_ind; intros. + apply Int63.leb_spec. simplify; auto. + pose proof (not_int63_lt_0 (PArray.length a)). + rewrite Int63.to_Z_0. auto. + split; intros. simpl in H. rewrite Int63.to_Z_0 in H. + rewrite <- Int63.to_Z_0 in H. apply Int63.ltb_spec in H. + now apply Misc.ltb_0 in H. auto. + split; [|erewrite Misc.to_Z_add_1 by eassumption; + rewrite Zlength_cons; lia]. + cbn [rev]. intros. erewrite Misc.to_Z_add_1 in H2 by eassumption. + assert (Int63.to_Z i0 < Int63.to_Z i \/ Int63.to_Z i0 = Int63.to_Z i) by lia. + inv H3. rewrite app_nth1; auto. apply H1; lia. apply Nat2Z.inj_lt. + rewrite rev_length. + rewrite <- Zlength_correct. + rewrite Z_to_nat_max. + apply Int63.leb_spec in H. + apply Int63.ltb_spec in H0. + pose proof (not_int63_lt_0 i0). lia. + apply Int63.leb_spec in H. + apply Int63.ltb_spec in H0. simpl in H. + rewrite app_nth2; rewrite rev_length. + replace (Datatypes.length a0) with (Z.to_nat (Z.of_nat (Datatypes.length a0))) + by (now apply Nat2Z.id). + rewrite <- Zlength_correct. + rewrite <- Z2Nat.inj_sub. rewrite Zeq_minus by lia. + simpl. apply Int63.to_Z_inj in H4. + now subst. + rewrite Zlength_correct. lia. + zify. rewrite <- Zlength_correct. lia. +Qed. + +Lemma to_list_corr : + forall A (a: PArray.array A) i, + Int63.to_Z i < Int63.to_Z (PArray.length a) -> + PArray.get a i = nth (Z.to_nat (Int63.to_Z i)) (Misc.to_list a) (PArray.default a). +Proof. + intros; pose proof (to_list_corr' A a). + inv H0; auto. +Qed. + +Lemma check_list_correct : + forall A comp (a: list A) b i da db + (CHK : is_true (check_list comp a b)) + (LT : (i < (length a))%nat), + is_true (comp (nth i b db) (nth i a da)). +Proof. + unfold is_true; induction a; try solve[crush]. + intros. + simpl in CHK. destruct_match; try discriminate. + apply andb_prop in CHK. inv CHK. + simplify. destruct i. auto. + apply IHa; auto. lia. +Qed. + +Lemma Zlength_aux_concat: forall A (l1 l2: list A) (acc: Z), + Zlength_aux acc _ (l1 ++ l2) = Zlength_aux (Zlength_aux acc _ l1) _ l2. +Proof. + induction l1 ; intros; eauto. + simpl. + rewrite IHl1; eauto. +Qed. + +Lemma length_to_list : + forall A (a: PArray.array A), + Int63.to_Z (PArray.length a) = Zlength (Misc.to_list a). +Proof. + intros. unfold Misc.to_list. + eapply Misc.foldi_ind; eauto. + - apply Misc.leb_0. + - intros. unfold Zlength in *. simpl. + rewrite Zlength_aux_concat. + rewrite <- H1. simpl. + erewrite Misc.to_Z_add_1; eauto. +Qed. + +Lemma check_list_length : + forall A comp (l1: list A) l2, + is_true (check_list comp l1 l2) -> + Zlength l1 <= Zlength l2. +Proof. + induction l1; crush. rewrite Zlength_nil. rewrite Zlength_correct. lia. + destruct l2; crush. rewrite !Zlength_cons. + unfold is_true in H. simplify. eapply IHl1 in H1. lia. +Qed. + +Lemma check_array_length : + forall A comp (a: PArray.array A) b, + is_true (check_array comp a b) -> + Int63.to_Z (PArray.length a) <= Int63.to_Z (PArray.length b). +Proof. + intros. unfold check_array in *. + eapply check_list_length in H. rewrite !length_to_list; auto. +Qed. + +Lemma check_array_correct : + forall A comp (a: PArray.array A) b, + is_true (check_array comp a b) -> + forall i, + Int63.to_Z i < Int63.to_Z (PArray.length a) -> + is_true (comp (PArray.get b i) (PArray.get a i)). +Proof. + unfold check_array, is_true; intros * CHK i LT. + pose proof (not_int63_lt_0 i). + repeat rewrite to_list_corr. + apply check_list_correct; auto. zify. rewrite <- Zlength_correct. + rewrite <- length_to_list. lia. + lia. apply check_array_length in CHK. lia. +Qed. + +Section PROOF. + + Context (t_i: PArray.array SMT_classes.typ_compdec). + Context (t_func: PArray.array (tval t_i)). + + Definition bo_dec: + forall (a b: option bool), + {a = b} + {a <> b}. + Proof. pose proof bool_dec. decide equality. Qed. + + Lemma of_list'_length : forall A, + forall (b: list A) acc array, + PArray.length (of_list' acc array b) = PArray.length array. + Proof. + induction b; intros. + - reflexivity. + - simpl of_list'. + erewrite IHb; eauto. + rewrite PArray.length_set. auto. + Qed. + + Lemma of_list_Zlength : forall A (d: A) b, + Zlength b < wB -> + PArray.length (of_list d b) = of_Z (Z.of_nat (length b)). + Proof. + unfold of_list. + intros. rewrite of_list'_length; eauto. + rewrite PArray.length_make. + unfold PArray.max_length. + destruct_match. + - rewrite Zlength_correct. auto. + - rewrite Misc.leb_negb_gtb in Heqb0. + rewrite negb_false_iff in Heqb0. + eapply ltb_spec in Heqb0. + rewrite of_Z_spec in Heqb0. + rewrite Z.mod_small in Heqb0; try lia. + + rewrite Misc.max_int_wB in Heqb0. lia. + + split; auto. rewrite Zlength_correct. lia. + Qed. + + Opaque PArray.set. + + Lemma Zlength_aux_pos : forall A (b: list A) acc, + acc >= 0 -> + Zlength_aux acc A b >= 0. + Proof. + induction b; intros. + - auto. + - simpl. eapply IHb; eauto. + lia. + Qed. + + Lemma Zlength_pos : forall A (b: list A), + Zlength b >= 0. + Proof. + unfold Zlength. + intros. eapply Zlength_aux_pos; eauto. + lia. + Qed. + + Lemma of_list'_idx_preserve : forall A (b: list A) idx array, + forall i, + (0 <= idx)%Z -> + (i + (0 <= idx + (Zlength b) < wB)%Z -> + (PArray.get (of_list' idx array b) i) = PArray.get array i. + Proof. + induction b; intros. + - auto. + - simpl of_list'. + rewrite Zlength_cons in *. + erewrite IHb; eauto. + + rewrite PArray.get_set_other; eauto. + intro ; subst. eelim Misc.not_ltb_refl; eauto. + + lia. + + rewrite ltb_spec in *. rewrite of_Z_spec in *. + assert (Zlength b >= 0) by (eapply Zlength_pos; eauto). + rewrite ! Z.mod_small in *; try lia. + + lia. + Qed. + + Opaque PArray.set. + + Lemma to_list_of_list'3 : + forall A (l: list A) arr x y, + y + Zlength l < wB -> + to_Z x < y -> + to_Z x < to_Z (PArray.length arr) -> + PArray.get (of_list' y arr l) x = PArray.get arr x. + Proof. + induction l; crush. + erewrite IHl. rewrite PArray.get_set_other; auto. + - unfold not; intros. assert (forall x y, x = y -> to_Z x = to_Z y) by (intros; now subst). + apply H3 in H2. rewrite of_Z_spec in H2. rewrite Z.mod_small in H2; subst. lia. + pose proof (not_int63_lt_0 x). split. lia. rewrite Zlength_cons in H. + pose proof (Zlength_pos _ l). lia. + - rewrite Zlength_cons in H. lia. + - lia. + - rewrite PArray.length_set. auto. + Qed. + + Lemma to_list_of_list'2 : + forall A l (d: A) x0 v a + (PARRR: to_Z x0 < to_Z (PArray.length a)) + (XXXXX: v + Zlength l < wB), + PArray.default a = d -> + to_Z (PArray.length a) <= Zlength l + v -> + v <= to_Z x0 -> + PArray.get (of_list' v a l) x0 = nth ((Z.to_nat (to_Z x0 - v))) l d. + Proof. + induction l; intros. + - simpl. assert ((x0 Z.of_nat x = Z.of_nat y) by (intros; subst; auto). + apply H2 in Heqn. rewrite Z_to_nat_max in Heqn. + rewrite Zlength_cons in H0. unfold Z.succ in H0. + { replace (Z.max (to_Z x0 - v) 0) with (to_Z x0 - v) in Heqn by lia. + assert (to_Z x0 = v) by lia. subst. + rewrite of_to_Z. erewrite to_list_of_list'3; try lia. + rewrite PArray.get_set_same; auto. + apply Int63.ltb_spec. auto. + rewrite Zlength_cons in XXXXX. lia. + rewrite PArray.length_set; auto. + } + rewrite IHl with (d:=d). + assert (Z.to_nat (to_Z x0 - (v + 1)) = n). + { assert (forall a b, a = b -> (a - 1)%nat = (b - 1)%nat) by lia. + apply H2 in Heqn. + assert ((S n - 1)%nat = n) by lia. rewrite H3 in Heqn. rewrite <- Heqn. + lia. + } rewrite H2. auto. + rewrite PArray.length_set. lia. rewrite Zlength_cons in XXXXX. lia. + rewrite PArray.default_set. auto. + rewrite PArray.length_set. rewrite Zlength_cons in H0. lia. + lia. + Qed. + Transparent PArray.set. + + Lemma to_list_of_list' : + forall A l (d: A) x0, + Zlength l < wB -> + to_Z x0 < Zlength l -> + PArray.get (of_list d l) x0 = nth (Z.to_nat (Int63.to_Z x0)) l d. + Proof. + unfold of_list. intros. + replace (to_Z x0) with (to_Z x0 - 0) by lia. + pose proof (Zlength_pos _ l). + apply to_list_of_list'2. + rewrite PArray.length_make. + assert ((of_Z (Zlength l) <=? PArray.max_length)%int63 = true). + apply Int63.leb_spec. rewrite of_Z_spec. rewrite Z.mod_small by lia. + unfold PArray.max_length. rewrite Misc.max_int_wB. lia. rewrite H2. + rewrite of_Z_spec. rewrite Z.mod_small by lia. auto. + lia. + rewrite PArray.default_make; auto. + rewrite PArray.length_make. + assert ((of_Z (Zlength l) <=? PArray.max_length)%int63 = true). + apply Int63.leb_spec. rewrite of_Z_spec. rewrite Z.mod_small by lia. + unfold PArray.max_length. rewrite Misc.max_int_wB. lia. rewrite H2. + rewrite of_Z_spec. rewrite Z.mod_small by lia. auto. lia. + pose proof (not_int63_lt_0 x0). lia. + Qed. + + Lemma to_list_of_list2 : (* lemmes nth of_list' *) + forall A (d: A) (b: list A) x, + Zlength b < wB -> + In x (Misc.to_list (of_list d b)) -> + In x b. + Proof. + intros. apply Misc.In_to_list in H0. simplify. + rewrite to_list_of_list'. apply nth_In. + apply Int63.ltb_spec in H0. rewrite of_list_Zlength in H0 by auto. + rewrite of_Z_spec in H0. rewrite Z.mod_small in H0. + assert (Z.of_nat (Z.to_nat φ (x0)%int63) < Z.of_nat (Datatypes.length b)). + { rewrite Z_to_nat_max. + pose proof (not_int63_lt_0 x0). + lia. + } + lia. + simplify. lia. + rewrite <- Zlength_correct; auto. + auto. + apply Int63.ltb_spec in H0. rewrite of_list_Zlength in H0; auto. + rewrite <- Zlength_correct in H0. rewrite Int63.of_Z_spec in H0. + rewrite Z.mod_small in H0; auto. + pose proof (Zlength_pos _ b). lia. + Qed. + + Lemma of_pos_rec_unfold_xO : forall n p, + of_pos_rec (S n) (xO p) = (of_pos_rec n p << 1)%int63. + Proof. + intros. reflexivity. + Qed. + + Lemma is_even_of_pos : + forall p a, + (Z.pos a - 2) * 2 = Z.pos p -> + is_even (of_pos p) = true. + Proof. + destruct p. + - lia. + - intros. + unfold of_pos, size. + rewrite of_pos_rec_unfold_xO. + eapply is_even_lsl_1. + - lia. + Qed. + + Lemma is_pos_P_of_P: + forall a, + 0 <= (Zpos a - 2) * 2 < wB -> + State.Lit.is_pos (P_of_P a) = true. + Proof. + unfold State.Lit.is_pos, P_of_P. + intros. + unfold of_Z. case_eq ((Z.pos a - 2) * 2); intros; try lia. + - apply Bva_checker.is_even_0. + - eapply is_even_of_pos; eauto. + Qed. + + Lemma of_pos_rec_unfold_xI : forall n p, + of_pos_rec (S n) (xI p) = (of_pos_rec n p << 1 lor 1)%int63. + Proof. + intros. reflexivity. + Qed. + + Lemma is_even_of_pos_false : + forall p a, + (Z.pos a - 2) * 2 + 1 = Z.pos p -> + is_even (of_pos p) = false. + Proof. + destruct p. + - intros. unfold of_pos, size. + rewrite of_pos_rec_unfold_xI. + rewrite Misc.is_even_or. + rewrite andb_false_r; auto. + - lia. + - intros. auto. + Qed. + + Lemma is_neg_N_of_P: + forall a, + 0 <= (Zpos a - 2) * 2 + 1 < wB -> + State.Lit.is_pos (N_of_P a) = false. + Proof. + unfold State.Lit.is_pos, N_of_P. + intros. + unfold of_Z. case_eq ((Z.pos a - 2) * 2 + 1); intros; try lia. + eapply is_even_of_pos_false; eauto. + Qed. + + Lemma blit_N_of_P: + forall a, + 1 < Z.pos a -> + (Z.pos a - 2) * 2 + 1 < wB -> + State.Lit.blit (N_of_P a) = of_P a. + Proof. + unfold State.Lit.blit, of_P, N_of_P; intros. + apply to_Z_inj. + rewrite lsr_spec. + replace (to_Z (1)%int63) with 1 by crush. + replace (2 ^ 1) with 2 by lia. + rewrite !of_Z_spec. + rewrite !Z.mod_small by lia. + rewrite Z.div_add_l by lia. + assert (X: 1 / 2 = 0) by auto. + rewrite X. rewrite Z.add_0_r. auto. + Qed. + + Lemma blit_P_of_P: + forall a, + 1 < Z.pos a -> + (Z.pos a - 2) * 2 < wB -> + State.Lit.blit (P_of_P a) = of_P a. + Proof. + unfold State.Lit.blit, of_P, P_of_P; intros. + apply to_Z_inj. + rewrite lsr_spec. + replace (to_Z (1)%int63) with 1 by crush. + replace (2 ^ 1) with 2 by lia. + rewrite !of_Z_spec. + rewrite !Z.mod_small by lia. + rewrite Z_div_mult by lia. auto. + Qed. + + Lemma in_tree_lt_array : + forall d tree f x y, + tree_key_wf _ tree -> + of_tree d (PTree.map (fun _ : positive => lform_to_form) tree) = Some f -> + tree ! x = Some y -> + is_true (of_P x + PArray.default y = d. + Proof. + intros. unfold of_tree in *. destruct_match; try discriminate; simplify. + rewrite PArray_default_fold_set; auto. + Qed. + + Lemma hash_value_in : + forall max f ht h ht', + hash_value max f ht = (h, ht') -> + ht' ! h = Some f. + Proof. + unfold hash_value. intros. + destruct_match. + - inv H. eapply find_tree_correct; eauto. + - inv H. rewrite PTree.gss. auto. + Qed. + + Lemma AHhash_value_in : + forall max f ht h ht', + AH.hash_value max f ht = (h, ht') -> + ht' ! h = Some f. + Proof. + unfold AH.hash_value; intros. + destruct_match. + - inv H. eapply AH.find_tree_correct; eauto. + - inv H. rewrite PTree.gss. auto. + Qed. + + Lemma pos_of_tree_max_key : + forall A (d: A) h11 h6 x h, + tree_a_wf h11 -> + (max_key h6 <= max_key h11)%positive -> + h11 ! h = h6 ! h -> + AH.find_tree x h6 = Some h -> + 1 < Z.pos h < wB. + Proof. + intros. eapply AH.find_tree_Some in H2. + pose proof H2. rewrite <- H1 in H3. + eapply max_key_correct in H2. + unfold tree_a_wf in *. inversion H as [RANGE [PROP [SOME1 [SOME2 SOME3]]]]. + apply PROP in H3. lia. + Qed. + + Definition match_one {A} (f: PTree.t A) lf := + forall y z, lf ! y = Some z -> f ! y = Some z. + + Definition match_all {A} (f: PTree.t A) lf := + Forall (match_one f) lf. + +Lemma match_one_refl : + forall A x, @match_one A x x. +Proof. unfold match_one; auto. Qed. + +Instance match_one_Reflexive {A} : Reflexive (@match_one A). +Proof. unfold Reflexive. auto using match_one_refl. Qed. + +Lemma match_one_trans : + forall A x y z, @match_one A x y -> match_one y z -> match_one x z. +Proof. unfold match_one; auto. Qed. + +Instance match_one_Transitive {A} : Transitive (@match_one A). +Proof. unfold Transitive. eauto using match_one_trans. Qed. + + Lemma tree_a_wf_lt : + forall h x y, + tree_a_wf h -> + h ! x = Some y -> + 0 <= Z.pos x - 2 < wB. + Proof. + unfold tree_a_wf; intros. pose proof H0 as X. apply max_key_correct in H0. + unfold gt_1 in *. inv H. inv H2. apply H in X. + all: lia. + Qed. + + Section PROOF_MORE. + + Context (f: form_t) (a: atom_t) (f_t: PTree.t lform) (a_t: PTree.t atom). + + Context (CHKF: is_true (check_form f)). + Context (CHKA: is_true (check_atom a)). + Context (WFF: tree_f_wf f_t). + Context (WFA: tree_a_wf a_t). + + Context (GENF: of_tree Ftrue (PTree.map (fun _ => lform_to_form) f_t) = Some f). + Context (GENA: of_tree (Acop CO_xH) a_t = Some a). + + Context (la: PTree.t bool). + Context (la_arr: PArray.array (tval t_i_empty)). + + Context (MATCH_LA: forall i : positive, 0 <= Z.pos i - 2 < wB -> + PArray.get la_arr (of_P i) = t_func_mk (Zval la ! i)). + + Lemma match_one_hash_value : + forall max x h h' y y', + hash_value max y h = (y', h') -> + match_one x h' -> + match_one x h. + Proof. + unfold match_one. + intros. + eapply H0; eauto. + eapply hash_constant; eauto. + Qed. + + Lemma AHmatch_one_hash_value : + forall max x h h' y y', + AH.hash_value max y h = (y', h') -> + match_one x h' -> + match_one x h. + Proof. + unfold match_one. + intros. + eapply H0; eauto. + eapply AH.hash_constant; eauto. + Qed. + + Lemma gt_1_map : + gt_1 (PTree.map (fun _ : positive => lform_to_form) f_t). + Proof. + inversion WFF; unfold gt_1 in *; simplify. + rewrite PTree.gmap in H1. unfold option_map in *. + destruct (f_t ! x) eqn:?. eapply H0. eauto. easy. + Qed. + + Lemma gt_1_a : + gt_1 a_t. + Proof. inversion WFA; crush. Qed. + + Lemma in_tree_bounded : + forall t p y, + t ! p = Some y -> + match_one f_t t -> + 0 <= (Z.pos p - 2) * 2 + 1 < wB. + Proof. + intros. unfold match_one in *. apply H0 in H. + pose proof H. apply max_key_correct in H. + unfold tree_f_wf in WFF. inversion WFF as [X1 [X2 [X3 X4]]]. + unfold gt_1 in X3. + apply X3 in H1. + lia. + Qed. + + Lemma in_tree_bounded2 : + forall t p y, + t ! p = Some y -> + match_one f_t t -> + 0 <= (Z.pos p - 2) * 2 < wB. + Proof. + intros. unfold match_one in *. apply H0 in H. + pose proof H. apply max_key_correct in H. + unfold tree_f_wf in WFF. inversion WFF as [X1 [X2 [X3 X4]]]. + unfold gt_1 in X3. + apply X3 in H1. + lia. + Qed. + + Lemma in_tree_bounded3 : + forall t p y, + t ! p = Some y -> + match_one f_t t -> + 0 <= Z.pos p - 2 < wB. + Proof. + intros. unfold match_one in *. apply H0 in H. + pose proof H. apply max_key_correct in H. + unfold tree_f_wf in WFF. inversion WFF as [X1 [X2 [X3 X4]]]. + unfold gt_1 in X3. + apply X3 in H1. + lia. + Qed. + + Lemma in_tree_bounded4 : + forall t p y, + t ! p = Some y -> + match_one a_t t -> + 0 <= Z.pos p - 2 < wB. + Proof. + intros. unfold match_one in *. apply H0 in H. + pose proof H. apply max_key_correct in H. + unfold tree_a_wf in WFA. inversion WFA as [X1 [X2 [X3 X4]]]. + unfold gt_1 in X2. + apply X2 in H1. + lia. + Qed. + + Lemma tree_f_wf_tree_key_wf: forall f_t, + tree_f_wf f_t -> + tree_key_wf _ f_t. + Proof. + clear. intros. + inversion H as [X1 [X2 [X3 [X4 X5]]]]. + unfold tree_key_wf. + repeat split; try lia. + auto. + Qed. + + Lemma tree_a_wf_tree_key_wf: forall f_t, + tree_a_wf f_t -> + tree_key_wf _ f_t. + Proof. + clear. intros. + inversion H as [X1 [X2 [X3 [X4 X5]]]]. + unfold tree_key_wf. + repeat split; try lia. + auto. + Qed. + + Lemma tree_key_wf_map : forall A B f_t (f: positive -> A -> B), + tree_key_wf _ f_t -> + tree_key_wf _ (PTree.map f f_t). + Proof. + clear. intros. + inversion H as [X1 [X2 X3]]. + unfold tree_key_wf. + repeat split; try (erewrite <- max_key_map; eauto). + rewrite <- gt_1_map_iff; eauto. + Qed. + + #[local] Hint Resolve gt_1_map : to_smt. + #[local] Hint Resolve gt_1_a : to_smt. + #[local] Hint Resolve in_tree_bounded : to_smt. + #[local] Hint Resolve in_tree_bounded2 : to_smt. + #[local] Hint Resolve in_tree_bounded3 : to_smt. + #[local] Hint Resolve in_tree_bounded4 : to_smt. + + Lemma of_tree_of_P_2 : + forall t p v, + t ! p = Some v -> + match_one f_t t -> + PArray.get f (of_P p) = (lform_to_form v). + Proof. + unfold match_one; intros. + eapply of_tree_of_P; eauto with to_smt. + - eapply tree_key_wf_map; eauto. + eapply tree_f_wf_tree_key_wf; eauto. + - rewrite PTree.gmap. apply H0 in H. rewrite H. auto. + Qed. + + Lemma of_tree_of_P_2_hv : + forall t p v x h, + hash_value x v h = (p, t) -> + match_one f_t t -> + PArray.get f (of_P p) = (lform_to_form v). + Proof. + intros; apply hash_value_in in H; auto. + eapply of_tree_of_P_2; eassumption. + Qed. + + Lemma AHof_tree_of_P_2 : + forall t p v, + t ! p = Some v -> + match_one a_t t -> + PArray.get a (of_P p) = v. + Proof. + intros. pose proof H0. unfold match_one in H0; intros. + pose proof H. apply H0 in H. + eapply of_tree_of_P; eauto with to_smt. + eapply tree_a_wf_tree_key_wf; eauto. + Qed. + + Lemma AHof_tree_of_P_2_hv : + forall t p v x h, + AH.hash_value x v h = (p, t) -> + match_one a_t t -> + PArray.get a (of_P p) = v. + Proof. + unfold match_one; intros. apply AHhash_value_in in H. + pose proof H. apply H0 in H. + eapply of_tree_of_P; eauto with to_smt. + eapply tree_a_wf_tree_key_wf; eauto. + Qed. + + Lemma is_true_Atom_wf : + is_true (Atom.wf a). + Proof. + unfold is_true in *. unfold check_atom in CHKA. repeat (destruct_match; try discriminate; []); auto. + Qed. + + Lemma is_true_Form_wf : + is_true (wf f). + Proof. + unfold is_true in *. unfold check_form in CHKF. repeat (destruct_match; try discriminate; []); crush. + Qed. + + Lemma is_true_default_a : + PArray.default a = Acop CO_xH. + Proof. + unfold is_true in *. unfold check_atom in CHKA. repeat (destruct_match; try discriminate; []); auto. + Qed. + + Lemma gt_1_in : + forall h h1 y, + h ! h1 = Some y -> + match_one f_t h -> + 1 < Z.pos h1. + Proof. + intros. unfold match_one in H0. apply H0 in H. + inv WFF. unfold gt_1 in *. + inv H2. inv H4. inv H5. + apply H2 in H. lia. + Qed. + + Lemma gt_1_in2 : + forall h h1 y, + h ! h1 = Some y -> + match_one f_t h -> + (Z.pos h1 - 2) * 2 < wB. + Proof. + intros. unfold match_one in H0. apply H0 in H. + inv WFF. apply max_key_correct in H. lia. + Qed. + + Lemma gt_1_in3 : + forall h h1 y, + h ! h1 = Some y -> + match_one f_t h -> + (Z.pos h1 - 2) * 2 + 1 < wB. + Proof. + intros. unfold match_one in H0. apply H0 in H. + inv WFF. apply max_key_correct in H. lia. + Qed. + + #[local] Hint Resolve gt_1_in : to_smt. + #[local] Hint Resolve gt_1_in2: to_smt. + #[local] Hint Resolve gt_1_in3: to_smt. + + Lemma blit_lt_length : + forall h1 h y, + h ! h1 = Some y -> + match_one f_t h -> + is_true (State.Lit.blit (P_of_P h1) + match_one f_t h -> + is_true (State.Lit.blit (P_of_P h1) + match_one f_t h -> + is_true (State.Lit.blit (N_of_P h1) + match_one f_t h -> + is_true (State.Lit.blit (N_of_P h1) + match_one f_t h -> + 0 <= (Z.pos h1 - 2) * 2 < wB. + Proof. + intros. inversion WFF as [BOUNDS [MAXKEY [GT [S2 S3]]]]. unfold gt_1 in GT. + unfold match_one in *. apply H0 in H. pose proof H as X. apply GT in H. + apply max_key_correct in X. lia. + Qed. + + Lemma lt_pos_in2 : + forall h1 y h max h', + hash_value max y h' = (h1, h) -> + match_one f_t h -> + 0 <= (Z.pos h1 - 2) * 2 < wB. + Proof. + intros. eapply hash_value_in in H. eapply lt_pos_in; eauto. + Qed. + + #[local] Hint Resolve is_true_Atom_wf : to_smt. + #[local] Hint Resolve is_true_Form_wf : to_smt. + #[local] Hint Resolve is_true_default_a : to_smt. + #[local] Hint Resolve blit_lt_length2 : to_smt. + #[local] Hint Resolve lt_pos_in2 : to_smt. + #[local] Hint Resolve blit_lt_length_N2 : to_smt. + #[local] Hint Resolve tree_a_wf_tree_key_wf : to_smt. + #[local] Hint Resolve tree_f_wf_tree_key_wf : to_smt. + + #[local] Opaque Z.sub. + + Lemma to_smt_i_correct : + forall h p i, + match_one a_t h -> + to_smt_i h p = Some i -> + Atom.interp_aux t_i_empty la_arr (PArray.get (Atom.t_interp t_i_empty la_arr a)) (PArray.get a i) + = Bval _ Typ.TZ (Zval (eval_hash_pred p la)). + Proof. + intros * MATCH TO_SMT. + unfold to_smt_i in TO_SMT. + destruct_match; try discriminate; simplify. + unfold tree_a_wf in WFA. inversion WFA as [BOUND [PROP [AT1 [AT2 AT3]]]]. + replace (of_pos 1) with (of_P 3) by auto. + erewrite ! of_tree_of_P; try eassumption; try lia; simpl. + rewrite Atom.t_interp_wf by auto with to_smt. + erewrite ! of_tree_of_P; try eassumption; try lia; auto. + repeat split; auto; try lia. + repeat split; auto; try lia. + unfold tree_a_wf in WFA. inversion WFA as [BOUND [PROP [AT1 [AT2 AT3]]]]. + replace (of_pos 2) with (of_P 4) by auto. + erewrite ! of_tree_of_P; try eassumption; try lia; simpl. + rewrite Atom.t_interp_wf by auto with to_smt. + erewrite ! of_tree_of_P; try eassumption; try lia; simpl. + rewrite Atom.t_interp_wf by auto with to_smt. + erewrite ! of_tree_of_P; try eassumption; try lia; simpl. auto. + repeat split; auto; try lia. + repeat split; auto; try lia. + repeat split; auto; try lia. + + repeat (destruct_match; try discriminate; []). simplify. + apply AH.find_tree_correct in Heqo. + erewrite ! AHof_tree_of_P_2; try eassumption; try lia; simpl. + destruct_match. + rewrite MATCH_LA in Heqt. + destruct (la ! p0) eqn:?; [destruct b|]; simplify; unfold t_func_mk, Tval in Heqt; + symmetry in Heqt; inv Heqt; eapply Eqdep.EqdepTheory.inj_pair2 in H5; subst; auto. + lia. + Qed. + + Lemma interp_eval_pred : + forall h13 h11 h6 p0 i h8 h10 h9 h12, + match_one f_t h13 -> + match_one a_t h11 -> + match_one a_t h6 -> + eval_hash_pred p0 la <> Some true -> + to_smt_i h6 p0 = Some i -> + AH.hash_value 1%positive (Abop (BO_eq Typ.TZ) i (Int63.of_pos 1)) h8 = (h10, h11) -> + hash_value 1%positive (LFatom (of_P h10)) h9 = (h12, h13) -> + State.Lit.interp (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) + (N_of_P h12) = true. + Proof. + intros * M1 M2 M3 EVAL TOSMT HASH1 HASH2. unfold State.Lit.interp. + unfold State.Var.interp, interp_state_var. + pose proof HASH1 as IN1. eapply AHhash_value_in in IN1. + pose proof HASH2 as IN2. eapply hash_value_in in IN2. + rewrite is_neg_N_of_P by eauto with to_smt. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_N_of_P by eauto with to_smt. + erewrite ! of_tree_of_P_2 by eauto with to_smt; simpl. + unfold interp_form_hatom, interp_hatom. + rewrite Atom.t_interp_wf by eauto with to_smt. + erewrite ! AHof_tree_of_P_2 by eauto with to_smt; simpl. + rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite to_smt_i_correct by eauto with to_smt. + replace (of_pos 1) with (of_P 3) by auto. + inversion WFA as [X1 [X2 [X3 [X4 X5]]]]. + erewrite of_tree_of_P; + [idtac|idtac|eauto with to_smt|eauto with to_smt|eauto with to_smt]. + simpl. rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite of_tree_of_P; + [idtac|idtac|eauto with to_smt|eauto with to_smt|eauto with to_smt]. + simpl. + destruct (eval_hash_pred p0 la) eqn:?; [destruct b|]; try discriminate; auto. + eapply tree_a_wf_tree_key_wf; eauto. + eapply tree_a_wf_tree_key_wf; eauto. + Qed. + + Lemma eval_hash_pred_Pand_lt : + forall la p1 p2, + Zval (eval_hash_pred (p1 ∧ p2) la) = + if Zval (eval_hash_pred p1 la) + match_one f_t h9 -> + match_one a_t h8 -> + to_smt_i h (Pbase fa0) = Some i -> + to_smt_i h p1_1 = Some i0 -> + to_smt_i h p1_2 = Some i1 -> + eval_hash_pred (equiv (Pbase fa0) (p1_1 ∧ p1_2)) la = Some true -> + State.Lit.interp (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) + (P_of_P h1) = true. + Proof. + intros * SMTAND M1 M2 SMT1 SMT2 SMT3 EVAL. unfold State.Lit.interp. + unfold to_smt_and in SMTAND. repeat ldestruct_match; []. inversion SMTAND; clear SMTAND; subst. + assert (match_one f_t h15) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h13) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h11) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h0) by (eapply match_one_hash_value; eauto). + assert (match_one a_t h5) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h3) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h) by (eapply AHmatch_one_hash_value; eauto). + rewrite is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. unfold interp_state_var. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_P_of_P. + erewrite of_tree_of_P_2_hv; try eassumption. simpl. + unfold State.Lit.interp, State.Var.interp. + rewrite !is_pos_P_of_P by eauto with to_smt. + rewrite !t_interp_wf by eauto with to_smt. + rewrite !blit_P_of_P. + erewrite !of_tree_of_P_2_hv; try eassumption; simpl. + unfold interp_form_hatom, interp_hatom. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !AHof_tree_of_P_2_hv; try eassumption; simpl. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !to_smt_i_correct by eassumption. simpl. + apply eval_equiv in EVAL. + replace (eval_hash_pred (Pbase fa0) la) with (la ! fa0) in EVAL by auto. + rewrite EVAL. + rewrite !eval_hash_pred_Pand_lt. + destruct_match; apply Typ.i_eqb_refl. + all: match goal with + | H: hash_value _ _ _ = (?v, _) |- _ < Z.pos ?v => + apply hash_value_in in H; eauto with to_smt + | H: hash_value _ _ _ = (?v, _) |- (Z.pos ?v - _) * _ < _ => + apply hash_value_in in H; eauto with to_smt + end. + Qed. + + Lemma to_smt_and_match_one : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_and h0 h i i0 i1 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one f_t h0. + Proof. + unfold to_smt_and; intros * SMTAND M1. + repeat ldestruct_match; []. inversion SMTAND; clear SMTAND; subst. + repeat (eapply match_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_and_match_one2 : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_and h0 h i i0 i1 = (h9, h8, h1) -> + match_one a_t h8 -> + match_one a_t h. + Proof. + unfold to_smt_and; intros * SMTAND M1. + repeat ldestruct_match; []. inversion SMTAND; clear SMTAND; subst. + repeat (eapply AHmatch_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_or_correct : + forall h fa0 i p1_1 i0 p1_2 i1 h0 h9 h8 h1, + to_smt_or h0 h i i0 i1 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one a_t h8 -> + to_smt_i h (Pbase fa0) = Some i -> + to_smt_i h p1_1 = Some i0 -> + to_smt_i h p1_2 = Some i1 -> + eval_hash_pred (equiv (Pbase fa0) (p1_1 ∨ p1_2)) la = Some true -> + State.Lit.interp (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) + (P_of_P h1) = true. + Proof. + intros * SMTOR M1 M2 SMT1 SMT2 SMT3 EVAL. unfold State.Lit.interp. + unfold to_smt_or in SMTOR. repeat ldestruct_match; []. inversion SMTOR; clear SMTOR; subst. + assert (match_one f_t h15) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h13) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h11) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h0) by (eapply match_one_hash_value; eauto). + assert (match_one a_t h5) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h3) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h) by (eapply AHmatch_one_hash_value; eauto). + rewrite is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. unfold interp_state_var. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_P_of_P. + erewrite of_tree_of_P_2_hv; try eassumption. simpl. + unfold State.Lit.interp, State.Var.interp. + rewrite !is_pos_P_of_P by eauto with to_smt. + rewrite !t_interp_wf by eauto with to_smt. + rewrite !blit_P_of_P. + erewrite !of_tree_of_P_2_hv; try eassumption; simpl. + unfold interp_form_hatom, interp_hatom. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !AHof_tree_of_P_2_hv; try eassumption; simpl. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !to_smt_i_correct by eassumption. simpl. + apply eval_equiv in EVAL. + replace (eval_hash_pred (Pbase fa0) la) with (la ! fa0) in EVAL by auto. + rewrite EVAL. + rewrite !eval_hash_pred_Por_lt. + destruct_match; apply Typ.i_eqb_refl. + all: match goal with + | H: hash_value _ _ _ = (?v, _) |- _ < Z.pos ?v => + apply hash_value_in in H; eauto with to_smt + | H: hash_value _ _ _ = (?v, _) |- (Z.pos ?v - _) * _ < _ => + apply hash_value_in in H; eauto with to_smt + end. + Qed. + + Lemma to_smt_or_match_one : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_or h0 h i i0 i1 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one f_t h0. + Proof. + unfold to_smt_or; intros * SMTOR M1. + repeat ldestruct_match; []. inversion SMTOR; clear SMTOR; subst. + repeat (eapply match_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_or_match_one2 : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_or h0 h i i0 i1 = (h9, h8, h1) -> + match_one a_t h8 -> + match_one a_t h. + Proof. + unfold to_smt_or; intros * SMTOR M1. + repeat ldestruct_match; []. inversion SMTOR; clear SMTOR; subst. + repeat (eapply AHmatch_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_imp_correct : + forall h fa0 i p1_1 i0 p1_2 i1 h0 h9 h8 h1, + to_smt_imp h0 h i i0 i1 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one a_t h8 -> + to_smt_i h (Pbase fa0) = Some i -> + to_smt_i h p1_1 = Some i0 -> + to_smt_i h p1_2 = Some i1 -> + eval_hash_pred (equiv (Pbase fa0) (p1_1 → p1_2)) la = Some true -> + State.Lit.interp (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) + (P_of_P h1) = true. + Proof. + intros * SMTIMP M1 M2 SMT1 SMT2 SMT3 EVAL. unfold State.Lit.interp. + unfold to_smt_imp in SMTIMP. repeat ldestruct_match; []. inversion SMTIMP; clear SMTIMP; subst. + assert (match_one f_t h19) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h17) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h15) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h0) by (eapply match_one_hash_value; eauto). + assert (match_one a_t h11) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h7) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h5) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h3) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h) by (eapply AHmatch_one_hash_value; eauto). + rewrite is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. unfold interp_state_var. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_P_of_P. + erewrite of_tree_of_P_2_hv by eassumption. simpl. + unfold State.Lit.interp, State.Var.interp. + rewrite !is_pos_P_of_P by eauto with to_smt. + rewrite !t_interp_wf by eauto with to_smt. + rewrite !blit_P_of_P. + erewrite !of_tree_of_P_2_hv by eassumption; simpl. + unfold interp_form_hatom, interp_hatom. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !AHof_tree_of_P_2_hv by eassumption; simpl. + rewrite !Atom.t_interp_wf by auto with to_smt. + inversion WFA as [WFA1 [WFA2 [WFA3 [WFA4 WFA5]]]]. + erewrite !of_tree_of_P by + (try eassumption; try (eapply tree_a_wf_tree_key_wf; eauto); lia). + simpl. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !of_tree_of_P by + (try eassumption; try (eapply tree_a_wf_tree_key_wf; eauto); lia). + erewrite !AHof_tree_of_P_2_hv by eassumption. + erewrite !to_smt_i_correct by eassumption. cbn [Atom.interp_aux]. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !AHof_tree_of_P_2_hv by eassumption. + erewrite !to_smt_i_correct by eassumption. cbn [Atom.interp_aux]. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !of_tree_of_P by + (try eassumption; try (eapply tree_a_wf_tree_key_wf; eauto); lia). + cbn [Atom.interp_aux]. + rewrite !Atom.t_interp_wf by auto with to_smt. + erewrite !of_tree_of_P by + (try eassumption; try (eapply tree_a_wf_tree_key_wf; eauto); lia). + erewrite !to_smt_i_correct by eassumption. + cbn -[Z.sub]. + apply eval_equiv in EVAL. + replace (eval_hash_pred (Pbase fa0) la) with (la ! fa0) in EVAL by auto. + rewrite EVAL. + rewrite !eval_hash_pred_Pimp_lt. + destruct_match; auto using Z.eqb_refl. + all: match goal with + | H: hash_value _ _ _ = (?v, _) |- _ < Z.pos ?v => + apply hash_value_in in H; eauto with to_smt + | H: hash_value _ _ _ = (?v, _) |- (Z.pos ?v - _) * _ < _ => + apply hash_value_in in H; eauto with to_smt + end. + Qed. + + Lemma to_smt_imp_match_one : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_imp h0 h i i0 i1 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one f_t h0. + Proof. + unfold to_smt_imp; intros * SMTIMP M1. + repeat ldestruct_match; []. inversion SMTIMP; clear SMTIMP; subst. + repeat (eapply match_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_imp_match_one2 : + forall h i i0 i1 h0 h9 h8 h1, + to_smt_imp h0 h i i0 i1 = (h9, h8, h1) -> + match_one a_t h8 -> + match_one a_t h. + Proof. + unfold to_smt_imp; intros * SMTIMP M1. + repeat ldestruct_match; []. inversion SMTIMP; clear SMTIMP; subst. + repeat (eapply AHmatch_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_not_correct : + forall h fa0 i p1_1 i0 h0 h9 h8 h1, + to_smt_not h0 h i i0 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one a_t h8 -> + to_smt_i h (Pbase fa0) = Some i -> + to_smt_i h p1_1 = Some i0 -> + eval_hash_pred (equiv (Pbase fa0) (¬ p1_1)) la = Some true -> + State.Lit.interp (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) + (P_of_P h1) = true. + Proof. + intros * SMTNOT M1 M2 SMT1 SMT2 EVAL. unfold State.Lit.interp. + unfold to_smt_not in SMTNOT. repeat ldestruct_match; []. inversion SMTNOT; clear SMTNOT; subst. + assert (match_one f_t h0) by (eapply match_one_hash_value; eauto). + assert (match_one a_t h3) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h) by (eapply AHmatch_one_hash_value; eauto). + rewrite is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. unfold interp_state_var. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_P_of_P. + erewrite of_tree_of_P_2_hv by eassumption. cbn. + unfold interp_form_hatom. unfold interp_hatom. + rewrite Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2_hv by eassumption. cbn. + rewrite Atom.t_interp_wf by eauto with to_smt. + inversion WFA as [WFA1 [WFA2 [WFA3 [WFA4 WFA5]]]]. + erewrite !to_smt_i_correct by eassumption. + rewrite Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2_hv by eassumption. cbn. + rewrite Atom.t_interp_wf by eauto with to_smt. + erewrite !to_smt_i_correct by eassumption. cbn. + apply eval_equiv in EVAL. + replace (eval_hash_pred (Pbase fa0) la) with (la ! fa0) in EVAL by auto. + rewrite EVAL. + rewrite !eval_hash_pred_Pnot_lt. + auto using Z.eqb_refl. + all: match goal with + | H: hash_value _ _ _ = (?v, _) |- _ < Z.pos ?v => + apply hash_value_in in H; eauto with to_smt + | H: hash_value _ _ _ = (?v, _) |- (Z.pos ?v - _) * _ < _ => + apply hash_value_in in H; eauto with to_smt + end. + Qed. + + Lemma to_smt_not_match_one : + forall h i i0 h0 h9 h8 h1, + to_smt_not h0 h i i0 = (h9, h8, h1) -> + match_one f_t h9 -> + match_one f_t h0. + Proof. + unfold to_smt_not; intros * SMTNOT M1. + repeat ldestruct_match; []. inversion SMTNOT; clear SMTNOT; subst. + repeat (eapply match_one_hash_value; try eassumption). + Qed. + + Lemma to_smt_not_match_one2 : + forall h i i0 h0 h9 h8 h1, + to_smt_not h0 h i i0 = (h9, h8, h1) -> + match_one a_t h8 -> + match_one a_t h. + Proof. + unfold to_smt_not; intros * SMTNOT M1. + repeat ldestruct_match; []. inversion SMTNOT; clear SMTNOT; subst. + repeat (eapply AHmatch_one_hash_value; try eassumption). + Qed. + + Lemma fold_to_smt_l_correct : + forall l l0 h7 h6 l1 h9 h8 x, + match_one f_t h9 -> match_one a_t h8 -> + fold_right to_smt_l (Some (l0, h7, h6)) l = Some (l1, h9, h8) -> + (forall y, In y l0 -> State.Lit.interp (interp_state_var + (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) y = true) -> + eval_hash_pred (fold_left Pand (to_equiv l) T) la = Some true -> + In x l1 -> + State.Lit.interp (interp_state_var + (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) x = true. + Proof. + induction l as [|a0 l' IHl]; intros * M1 M2 TO_SMT_L PREV_LIST EVAL_FOLD IN. simplify; eauto. + simpl in EVAL_FOLD. + apply eval_hash_pred_fold_Pand4 in EVAL_FOLD. inversion EVAL_FOLD as [EVAL_FOLD' EVAL_A0]. + destruct a0 as [fa0 sa0]. cbn [fst snd] in *. + simpl in TO_SMT_L. unfold to_smt_l in TO_SMT_L. fold to_smt_l in TO_SMT_L. + repeat (destruct_match; try discriminate; []). + destruct_match; try discriminate; + repeat (destruct_match; try discriminate; []); inversion TO_SMT_L; subst; + [ assert (match_one f_t h0) by (eapply to_smt_and_match_one; eauto); + assert (match_one a_t h) by (eapply to_smt_and_match_one2; eauto) + | assert (match_one f_t h0) by (eapply to_smt_or_match_one; eauto); + assert (match_one a_t h) by (eapply to_smt_or_match_one2; eauto) + | assert (match_one f_t h0) by (eapply to_smt_imp_match_one; eauto); + assert (match_one a_t h) by (eapply to_smt_imp_match_one2; eauto) + | assert (match_one f_t h0) by (eapply to_smt_not_match_one; eauto); + assert (match_one a_t h) by (eapply to_smt_not_match_one2; eauto) ]. + all: apply in_app_or in IN; inversion IN as [INL | INP]; clear IN; [solve [eauto]|]; + inversion INP as [EQ | NIL]; clear INP; [|solve [inversion NIL]]; subst. + - eapply to_smt_and_correct; eauto. now rewrite <- eval_hash_pred_T_Pand. + - eapply to_smt_or_correct; eauto. now rewrite <- eval_hash_pred_T_Pand. + - eapply to_smt_imp_correct; eauto. now rewrite <- eval_hash_pred_T_Pand. + - eapply to_smt_not_correct; eauto. now rewrite <- eval_hash_pred_T_Pand. + Qed. + + Lemma gt_1_fold_init : + forall (lx: list unit) v' l0 init_forms init_atoms v l1 form_defs atom_defs, + fold_right (fun _ => declare_atoms_with_bounds) + (v', l0, init_forms, init_atoms) + lx = (v, l1, form_defs, atom_defs) -> + 0 <= v' -> 0 <= v. + Proof. + induction lx; crush. + remember (fold_right (fun _ : unit => declare_atoms_with_bounds) + (v', l0, init_forms, init_atoms) lx) as folds. + symmetry in Heqfolds. destruct folds as [[[A1 A2] A3] A4]. + eapply IHlx in Heqfolds; eauto. unfold declare_atoms_with_bounds in H. + repeat (destruct_match; []). simplify. lia. + Qed. + + Lemma fold_init_correct : + forall (lx: list unit) l1 x init_forms init_atoms form_defs atom_defs l0 v v', + match_one f_t form_defs -> match_one a_t atom_defs -> + fold_right (fun _ => declare_atoms_with_bounds) + (v', l0, init_forms, init_atoms) + lx = (v, l1, form_defs, atom_defs) -> + (forall y, In y l0 -> + State.Lit.interp (interp_state_var + (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) y = true) -> + In x l1 -> 0 <= v' -> v < wB -> + State.Lit.interp (interp_state_var + (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) x = true. + Proof. + induction lx; crush. assert (VPRIME := H4). clear H4. assert (VPRIME2 := H5). clear H5. + remember (fold_right (fun _ : unit => declare_atoms_with_bounds) (v', l0, init_forms, init_atoms) lx) as folds. + symmetry in Heqfolds. destruct folds as [[[V R] F] A]. simpl in H1. + repeat (destruct_match; []). simplify. + eapply in_app in H3. inv H3. eapply IHlx in Heqfolds; eauto; try lia. + assert (match_one f_t h4) by (eapply match_one_hash_value; eauto). + assert (match_one f_t h10) by (eapply match_one_hash_value; eauto). + eapply match_one_hash_value; eauto. + assert (match_one a_t h14) by (eapply AHmatch_one_hash_value; eauto). + assert (match_one a_t h0) by (eapply AHmatch_one_hash_value; eauto). + eapply AHmatch_one_hash_value; eauto. + assert (F1: match_one f_t h4) by (eapply match_one_hash_value; eauto). + assert (F2: match_one f_t h10) by (eapply match_one_hash_value; eauto). + assert (A1: match_one a_t h14) by (eapply AHmatch_one_hash_value; eauto). + assert (A2: match_one a_t h0) by (eapply AHmatch_one_hash_value; eauto). + inv H1; [|inv H3]. + unfold State.Lit.interp. + rewrite is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. unfold interp_state_var. + rewrite t_interp_wf by eauto with to_smt. + rewrite blit_P_of_P. + erewrite of_tree_of_P_2_hv by eassumption. cbn. + unfold State.Lit.interp. + rewrite !is_pos_P_of_P by eauto with to_smt. + unfold State.Var.interp. rewrite !t_interp_wf by eauto with to_smt. + rewrite !blit_P_of_P. + erewrite !of_tree_of_P_2_hv by eassumption. cbn. + unfold interp_form_hatom, interp_hatom. + rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2_hv by eassumption. cbn. + inversion WFA as [WFA1 [WFA2 [WFA3 [WFA4 WFA5]]]]. + rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2 by (try eassumption; reflexivity). cbn. + rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2 by (try eassumption; reflexivity). cbn. + rewrite !Atom.t_interp_wf by eauto with to_smt. + erewrite !AHof_tree_of_P_2 by (try eassumption; reflexivity). cbn. + erewrite !AHof_tree_of_P_2_hv by eassumption. cbn. + repeat (destruct_match; []). repeat (ldestruct_match; []). + assert (0 <= V). + { eapply gt_1_fold_init; eauto. } + replace (of_Z V) with (of_P (Z.to_pos (V + 2))) in Heqt by + (unfold of_P; f_equal; rewrite Z2Pos.id by lia; lia). + rewrite MATCH_LA in Heqt. + destruct (la ! (Z.to_pos (V + 2))). destruct b. + unfold t_func_mk, Tval in Heqt. cbn in Heqt. + symmetry in Heqt. inv Heqt. eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. + simpl in Heqb. unfold Bval in Heqb. symmetry in Heqb. inv Heqb. + eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. auto. + unfold t_func_mk, Tval in Heqt. cbn in Heqt. + symmetry in Heqt. inv Heqt. eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. + simpl in Heqb. unfold Bval in Heqb. symmetry in Heqb. inv Heqb. + eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. auto. + unfold t_func_mk, Tval in Heqt. cbn in Heqt. + symmetry in Heqt. inv Heqt. eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. + simpl in Heqb. unfold Bval in Heqb. symmetry in Heqb. inv Heqb. + eapply Eqdep.EqdepTheory.inj_pair2 in H5. subst. auto. + all: try match goal with + | H: hash_value _ _ _ = (?v, _) |- _ < Z.pos ?v => + apply hash_value_in in H; eauto with to_smt + | H: hash_value _ _ _ = (?v, _) |- (Z.pos ?v - _) * _ < _ => + apply hash_value_in in H; eauto with to_smt + end. + rewrite Z2Pos.id; lia. + Qed. + +End PROOF_MORE. + + Definition pl_leq {A} (a b: positive * A): Prop := + ((fst a) <= (fst b))%positive. + + Definition pl_lt {A} (a b: positive * A): Prop := + ((fst a) < (fst b))%positive. + + Definition pl_eq {A} (a b: positive * A): Prop := + fst a = fst b. + + Lemma pl_lt_trans : + forall A (a b c: positive * A), pl_lt a b -> pl_lt b c -> pl_lt a c. + Proof. unfold pl_lt; lia. Qed. + + Lemma pl_lt_irrefl : + forall A (a: positive * A), ~ pl_lt a a. + Proof. unfold pl_lt; lia. Qed. + + Lemma pl_lt_asymm : + forall A (a b: positive * A), pl_lt a b -> pl_lt b a -> False. + Proof. unfold pl_lt; lia. Qed. + + Definition Sorted := Sorted (@pl_lt predicate). + + Lemma pl_leq_refl : + forall A (a: positive * A), pl_leq a a. + Proof. unfold pl_leq; reflexivity. Qed. + + Lemma pl_leq_trans : + forall A (a b c: positive * A), pl_leq a b -> pl_leq b c -> pl_leq a c. + Proof. unfold pl_leq; intros; lia. Qed. + + Lemma pl_eq_refl : + forall A (a: positive * A), pl_eq a a. + Proof. unfold pl_leq; reflexivity. Qed. + + Lemma pl_eq_symm : + forall A (a b: positive * A), pl_eq a b -> pl_eq b a. + Proof. unfold pl_eq; intros; lia. Qed. + + Lemma pl_eq_trans : + forall A (a b c: positive * A), pl_eq a b -> pl_eq b c -> pl_eq a c. + Proof. unfold pl_eq; intros; lia. Qed. + + #[global] Instance Reflexive_pl_leq {A} : Reflexive (@pl_leq A). + Proof. unfold Reflexive. apply pl_leq_refl. Qed. + + #[global] Instance Transitive_pl_leq {A} : Transitive (@pl_leq A). + Proof. unfold Transitive. apply pl_leq_trans. Qed. + + #[global] Instance PreOrder_pl_leq {A} : PreOrder (@pl_leq A) := {}. + + #[global] Instance Reflexive_pl_eq {A} : Reflexive (@pl_eq A). + Proof. unfold Reflexive. apply pl_eq_refl. Qed. + + #[global] Instance Symmetric_pl_eq {A} : Symmetric (@pl_eq A). + Proof. unfold Symmetric. apply pl_eq_symm. Qed. + + #[global] Instance Transitive_pl_eq {A} : Transitive (@pl_eq A). + Proof. unfold Transitive. apply pl_eq_trans. Qed. + + #[global] Instance Equivalence_pl_eq {A} : Equivalence (@pl_eq A) := {}. + + #[global] Instance Antisymmetric_pl_leq {A} : Antisymmetric (positive * A) pl_eq pl_leq. + Proof. unfold Antisymmetric, pl_eq, pl_leq; lia. Qed. + + #[global] Instance PartialOrder_pl_leq {A} : PartialOrder (@pl_eq A) (@pl_leq A). + Proof. + unfold PartialOrder, relation_equivalence, predicate_equivalence. crush. + unfold relation_conjunction, predicate_intersection. simplify. unfold pl_eq, pl_leq, Basics.flip. + split. lia. lia. + Qed. + + #[global] Instance Irreflexive_pl_lt {A} : Irreflexive (@pl_lt A). + Proof. unfold Irreflexive, Reflexive, complement. apply pl_lt_irrefl. Qed. + + #[global] Instance Asymmetric_pl_lt {A} : Asymmetric (@pl_lt A). + Proof. unfold Asymmetric. apply pl_lt_asymm. Qed. + + #[global] Instance Transitive_pl_lt {A} : Transitive (@pl_lt A). + Proof. unfold Transitive. apply pl_lt_trans. Qed. + + #[global] Instance StrictOrder_pl_lt {A} : StrictOrder (@pl_lt A) := {}. + + Lemma eval_hash_pred_upd_asgn : + forall l p a, + Forall (fun x => (max_hash_pred p < fst x)%positive) l -> + eval_hash_pred p (upd_asgn l a) = eval_hash_pred p a. + Proof. + induction l as [| c l IHl]; try solve [crush]. + intros * FORALL. + inversion FORALL as [|? ? EL FORALL2]; subst; clear FORALL. + unfold upd_asgn. cbn [fold_left]. destruct_match; + replace ((fold_left + (fun (a' : PTree.t bool) (el : positive * hash_pred) => + match eval_hash_pred (snd el) a' with + | Some p_b => PTree.set (fst el) p_b a' + | None => a' + end) l a)) with (upd_asgn l a) in * by auto; eauto. + rewrite IHl; auto. + rewrite eval_hash_pred_gso; auto. + Qed. + + Lemma upd_asgn_not_in : + forall l c1 a, + ~ In c1 (map fst l) -> + (upd_asgn l a) ! c1 = a ! c1. + Proof. + induction l as [| c l IHl ]; try solve [crush]. + intros * NOTIN. simplify. + assert (fst c <> c1). + { unfold not; intros. subst. apply NOTIN; tauto. } + assert (~ In c1 (map fst l)). + { unfold not; intros. apply NOTIN; tauto. } + destruct_match; auto. + erewrite IHl; eauto. rewrite PTree.gso; auto. + Qed. + + Definition equiv_list_wf l := + Sorted l /\ Forall (fun x => (max_hash_pred (snd x) < fst x)%positive) l. + + Lemma equiv_list_wf_inv : + forall l a, + equiv_list_wf (a :: l) -> + equiv_list_wf l /\ Forall (pl_lt a) l /\ (max_hash_pred (snd a) < fst a)%positive. + Proof. + unfold equiv_list_wf; intros * S. inversion S as [A B]. + split; split. + - inversion A; auto. + - inversion B; auto. + - apply Sorted_StronglySorted in A. inversion A; auto. + unfold Relations_1.Transitive. apply pl_lt_trans. + - inversion B; auto. + Qed. + + Lemma equiv_list_wf_nil : + equiv_list_wf nil. + Proof. + unfold equiv_list_wf; split; auto; apply Sorted_nil. + Qed. + + Lemma sorted_not_in : + forall A l c1 (c2: A), + Forall (pl_lt (c1, c2)) l -> + ~ In c1 (map fst l). + Proof. + unfold not; intros. + eapply list_in_map_inv in H0. inv H0. inv H1. + eapply Forall_forall in H; eauto. + destruct x. simpl in H. unfold pl_lt in H. simpl in H. lia. + Qed. + + Lemma upd_asgn_set_sorted : + forall l a x, + Forall (Pos.lt x) (map fst l) -> + (upd_asgn l a) ! x = a ! x. + Proof. + induction l; auto. + intros. cbn. inv H. erewrite IHl; auto. + destruct_match; auto; now rewrite PTree.gso by lia. + Qed. + + Lemma eval_hash_pred_equiv_asgn : + forall l p a, + eval_hash_pred p a = Some true -> + equiv_list_wf l -> + Forall (fun x => a ! (fst x) = None) l -> + Forall (fun x => (max_hash_pred p < fst x)%positive) l -> + eval_hash_pred (fold_left Pand (to_equiv l) p) (upd_asgn l a) = Some true. + Proof. + induction l as [|c l IHl]; [crush|]. intros p a EVAL WF FORALL MAX. + destruct c as [c1 c2]. simpl. + inversion FORALL as [| ? ? NONE FORALL2]; subst; cbn [fst snd] in *; clear FORALL. + inversion MAX as [| ? ? MAXC MAX2]; subst; cbn [fst snd] in *; clear MAX. + eapply equiv_list_wf_inv in WF. inversion WF as [WF2 [FRL MHP]]; subst; clear WF. + cbn [snd fst] in *. destruct_match. + - rewrite IHl; eauto. + apply pand_true2; [|rewrite eval_hash_pred_gso by lia; auto]. + apply eval_equiv2. symmetry; rewrite eval_hash_pred_gso by crush. + rewrite Heqo. crush. symmetry; apply PTree.gss. + apply Forall_forall; intros. repeat match goal with H: Forall _ _ |- _ => eapply Forall_forall in H; eauto end. + unfold pl_lt in *. simplify. rewrite PTree.gso by lia. auto. + apply Forall_forall; intros. repeat match goal with H: Forall _ _ |- _ => eapply Forall_forall in H; eauto end. + unfold pl_lt in *. simplify. lia. + - eapply IHl; eauto. apply pand_true2. + apply eval_equiv2. rewrite Heqo. + simpl. auto. auto. + apply Forall_forall; intros. repeat match goal with H: Forall _ _ |- _ => eapply Forall_forall in H; eauto end. + unfold pl_lt in *. simplify. lia. + Qed. + +#[local] Opaque max_key. +#[local] Opaque Z.sub. +#[local] Opaque PTree.get. + +Lemma check_tree_f_wf_correct : + forall f, + check_tree_f_wf f = true -> tree_f_wf f. +Proof. + unfold check_tree_f_wf, tree_f_wf. + simplify; auto. + unfold gt_1; intros. + assert ((x = 1 \/ 1 < x)%positive) by lia. + inv H5; auto. rewrite H0 in H3. discriminate. + repeat (destruct_match; try discriminate; []); auto. + repeat (destruct_match; try discriminate; []); auto. +Qed. + +Lemma check_tree_a_wf_correct : + forall a, + check_tree_a_wf a = true -> tree_a_wf a. +Proof. + unfold check_tree_a_wf, tree_a_wf. + simplify; auto. + unfold gt_1; intros. + assert ((x = 1 \/ 1 < x)%positive) by lia. + inv H6; auto. rewrite H in H4. discriminate. + repeat (destruct_match; try discriminate; []); auto. + repeat (destruct_match; try discriminate; []); auto. + subst. apply Int63.eqb_spec in H2. apply Int63.eqb_spec in H1. subst. + auto. + repeat (destruct_match; try discriminate; []); auto. + subst. apply Int63.eqb_spec in H2. apply Int63.eqb_spec in H1. subst. + auto. +Qed. + +#[local] Transparent max_key. +#[local] Transparent Z.sub. +#[local] Transparent PTree.get. + +Lemma to_smt_l_match_one2 : + forall a y l0 h0 h r' f' a', + to_smt_l a (Some (l0, h0, h)) = Some (r', f', a') -> + match_one y a' -> + match_one y h. +Proof. + destruct a. destruct p0; simplify; try discriminate; + repeat (destruct_match; try discriminate; []); simplify; + eauto using to_smt_and_match_one2, to_smt_or_match_one2, + to_smt_imp_match_one2, to_smt_not_match_one2. +Qed. + +Lemma to_smt_l_match_one : + forall a y l0 h0 h r' f' a', + to_smt_l a (Some (l0, h0, h)) = Some (r', f', a') -> + match_one y f' -> + match_one y h0. +Proof. + destruct a. destruct p0; simplify; try discriminate; + repeat (destruct_match; try discriminate; []); simplify; + eauto using to_smt_and_match_one, to_smt_or_match_one, + to_smt_imp_match_one, to_smt_not_match_one. +Qed. + +Lemma to_smt_l_fold_match_one2 : + forall l y r f a r' f' a', + fold_right to_smt_l (Some (r, f, a)) l = Some (r', f', a') -> + match_one y a' -> + match_one y a. +Proof. + induction l; crush. remember (fold_right to_smt_l (Some (r, f, a0)) l) as folds. + destruct folds; [|crush]; []. + symmetry in Heqfolds. + destruct p. destruct p. eapply IHl; eauto. + eapply to_smt_l_match_one2; eauto. +Qed. + +Lemma to_smt_l_fold_match_one : + forall l y r f a r' f' a', + fold_right to_smt_l (Some (r, f, a)) l = Some (r', f', a') -> + match_one y f' -> + match_one y f. +Proof. + induction l; crush. remember (fold_right to_smt_l (Some (r, f, a0)) l) as folds. + destruct folds; [|crush]; []. + symmetry in Heqfolds. + destruct p. destruct p. eapply IHl; eauto. + eapply to_smt_l_match_one; eauto. +Qed. + +#[local] Opaque Z.sub. + +Lemma valid_to_smt : + forall p r f a la max la_arr, + ~ (is_true (Euf_Checker.valid la_arr a f r)) -> + asgn_transl_spec la la_arr -> + to_smt max p = Some (r, f, a) -> + eval_hash_pred (fold_left Pand (to_equiv (snd p)) T) la = Some true -> + eval_pred_list p la = Some true. +Proof. + intros. + destruct (bo_dec (eval_pred_list p la) (Some true)); auto. + exfalso. eapply H. inv H0. inv H4. assert (DEFAULT1:=H0). + assert (DEFAULT2:=H5). clear H5. clear H0. + unfold to_smt in H1. repeat (destruct_match; try discriminate; []). + simplify. + unfold is_true. + unfold eval_pred_list in n. cbn [fst snd] in *. + apply pand_false in n; auto; []. + unfold Euf_Checker.valid. + rewrite Misc.afold_left_and. + apply forallb_forall. intros. + eapply to_list_of_list2 in H4; [|lia]. + apply in_app_or in H4. + (* These are all easy properties that need a bit of extra checking code inside of to_smt. *) + assert (tree_a_wf h11) by (now apply check_tree_a_wf_correct). + assert (tree_f_wf h13) by (now apply check_tree_f_wf_correct). + assert (match_one h11 h8) by (now (eapply AHmatch_one_hash_value; eauto)). + assert (match_one h13 h9) by (now (eapply match_one_hash_value; eauto)). + inv H4. + - eapply fold_to_smt_l_correct; eauto. + (* This is a proof that all the variables will only be between -1 and 1, which are the only + values allowed by Zval anyways, so there can't be larger values in the context. *) + assert (forall y : Int63.int, + In y l0 -> + State.Lit.interp + (interp_state_var (interp_form_hatom t_i_empty la_arr a) + (interp_form_hatom_bv t_i_empty la_arr a) f) y = + true). + { intros. eapply fold_init_correct with (form_defs := h7) (atom_defs := h6). + auto. auto. eassumption. eassumption. auto. auto. eauto. + eapply to_smt_l_fold_match_one; eauto. eapply to_smt_l_fold_match_one2; eauto. + eapply Heqp5. inversion 1. auto. lia. lia. + } auto. + - inv H13; [|inv H4]. eapply interp_eval_pred with (h6 := h6) (h11 := h11) (h13 := h13); eauto; + try reflexivity; eapply to_smt_l_fold_match_one2; eauto. +Qed. + +#[local] Transparent Z.sub. + + Lemma upd_asgn_get : + forall l y x a, + (upd_asgn l a) ! y = None -> + (upd_asgn (l ++ (y, x) :: nil) a) ! y = eval_hash_pred x (upd_asgn l a). + Proof. + intros. + unfold upd_asgn; rewrite fold_left_app; cbn [fold_left]. + cbn [fst snd]; destruct_match; auto; apply PTree.gss. + Qed. + + Lemma list_app_cons : + forall A l1 l2 (a: A), + (l1 ++ a :: nil) ++ l2 = l1 ++ a :: l2. + Proof. induction l1; crush. Qed. + + Lemma upd_asgn_app : + forall l1 l2 a, + upd_asgn (l1 ++ l2) a = upd_asgn l2 (upd_asgn l1 a). + Proof. intros; unfold upd_asgn; now rewrite fold_left_app. Qed. + + Lemma upd_asgn_pred_In : + forall p a l1 a0, + ~ pred_In (fst a) p -> + eval_hash_pred p (upd_asgn (l1 ++ a :: nil) a0) = eval_hash_pred p (upd_asgn l1 a0). + Proof. + induction p; try solve [crush]. + + intros. simpl. rewrite upd_asgn_app. simpl. destruct_match; auto. + destruct (peq (fst a0) a); subst. + - exfalso; apply H. constructor. + - apply PTree.gso; auto. + + simplify. rewrite IHp2. rewrite IHp1; auto. + unfold not; intros; apply H. constructor; auto. + unfold not; intros; apply H. apply pred_In_Pand2; auto. + + simplify. rewrite IHp2. rewrite IHp1; auto. + unfold not; intros; apply H. constructor; auto. + unfold not; intros; apply H. apply pred_In_Por2; auto. + + simplify. rewrite IHp2. rewrite IHp1; auto. + unfold not; intros; apply H. constructor; auto. + unfold not; intros; apply H. apply pred_In_Pimp2; auto. + + simplify. rewrite IHp; auto. + unfold not; intros; apply H. constructor; auto. + Qed. + + Lemma upd_asgn_app1 : + forall l2 l1 p a, + Forall (fun x => ~ pred_In (fst x) p) l2 -> + eval_hash_pred p (upd_asgn (l1 ++ l2) a) = eval_hash_pred p (upd_asgn l1 a). + Proof. + induction l2. + - intros. now rewrite app_nil_r. + - intros. rewrite <- list_app_cons. inv H. + rewrite IHl2; auto. + rewrite upd_asgn_pred_In; auto. + Qed. + + Lemma eval_hash_pred_upd_asgn_set : + forall l2 p a a0 a0', + ~ pred_In a p -> + Forall (fun x => a <> (fst x) /\ ~ pred_In a (snd x)) l2 -> + (forall x, x <> a -> a0 ! x = a0' ! x) -> + eval_hash_pred p (upd_asgn l2 a0) = eval_hash_pred p (upd_asgn l2 a0'). + Proof. + induction l2; intros. + crush. eapply eval_hash_pred_except; eauto. + inversion H0 as [| ? ? [NEQ PRED] FRL ]; subst. + simpl. + assert (eval_hash_pred (snd a) a0' = eval_hash_pred (snd a) a1). + { eapply eval_hash_pred_except; eauto. intros. symmetry. eapply H1. auto. } + setoid_rewrite H2. + destruct_match; setoid_rewrite Heqo; eauto. + eapply IHl2; eauto; intros. + destruct (peq x (fst a)); subst. + now rewrite !PTree.gss. + rewrite !PTree.gso by auto; eauto. + Qed. + + Lemma eval_hash_pred_upd_asgn_set2 : + forall l2 p a a0 y, + ~ pred_In a p -> + Forall (fun x => a <> (fst x) /\ ~ pred_In a (snd x)) l2 -> + eval_hash_pred p (upd_asgn l2 (PTree.set a y a0)) + = eval_hash_pred p (upd_asgn l2 a0). + Proof. + intros; eapply eval_hash_pred_upd_asgn_set; eauto. + intros; now apply PTree.gso. + Qed. + + Lemma upd_asgn_app2 : + forall l1 l2 p a, + Forall (fun y => Forall (fun x => fst y <> fst x /\ ~ pred_In (fst y) (snd x)) l2) l1 -> + (forall x, In x l1 -> ~ pred_In (fst x) p) -> + eval_hash_pred p (upd_asgn (l1 ++ l2) a) = eval_hash_pred p (upd_asgn l2 a). + Proof. + induction l1; auto. + intros; simpl. + inv H. + rewrite IHl1; auto. + destruct_match; auto. + apply eval_hash_pred_upd_asgn_set2; auto. + Qed. + + Lemma sorted_rev : + forall A R l (b: A) a, + LocallySorted (Basics.flip R) (l ++ (b :: a :: nil)) -> + R a b. + Proof. + induction l; crush. + inv H. auto. inv H. + assert (forall A (a b: list A), nil = a ++ b -> a = nil /\ b = nil). + { induction a1; crush. } + apply H in H2. crush. + eapply IHl. rewrite <- H1. auto. + Qed. + + Lemma sorted_rev2 : + forall A R l (b: A) a, + LocallySorted (Basics.flip R) (l ++ b :: nil) -> + R a b -> + LocallySorted (Basics.flip R) (l ++ (b :: a :: nil)). + Proof. + induction l; crush. + constructor. constructor. auto. + inv H. + assert (forall A (a b: list A), nil = a ++ b -> a = nil /\ b = nil). + { induction a1; crush. } + apply H in H3. crush. + replace (a :: l ++ b :: a0 :: nil) with ((a :: l ++ b :: nil) ++ a0 :: nil). + rewrite <- H2. + replace ((a :: b0 :: l0) ++ a0 :: nil) with (a :: b0 :: (l0 ++ a0 :: nil)). + constructor; auto. + replace (b0 :: l0 ++ a0 :: nil) with ((b0 :: l0) ++ a0 :: nil). + rewrite H2. + replace ((l ++ b :: nil) ++ a0 :: nil) with (l ++ b :: a0 :: nil). + eapply IHl; eauto. rewrite <- H2. auto. + rewrite list_app_cons. auto. + rewrite app_comm_cons. auto. + rewrite app_comm_cons. auto. + rewrite <- app_comm_cons. + rewrite list_app_cons. auto. + Qed. + + Lemma sorted_rev3 : + forall A R l, + Sorted.LocallySorted R l -> + Sorted.LocallySorted (Basics.flip R) (@rev A l). + Proof. + induction l; crush. inv H. constructor. + inv H. crush. constructor. + crush. + replace ((rev l0 ++ b :: nil) ++ a :: nil) with (rev l0 ++ b :: a :: nil). + eapply sorted_rev2; eauto. + now rewrite list_app_cons. + Qed. + + Lemma sorted_rev4 : + forall A R l, + Sorted.Sorted R l -> + Sorted.Sorted (Basics.flip R) (@rev A l). + Proof. + intros. apply Sorted_LocallySorted_iff. + apply sorted_rev3. apply Sorted_LocallySorted_iff. auto. + Qed. + + Lemma unnest_predicate_lt_fresh : + forall p p' fresh fresh' lp, + unnest_predicate fresh p = (p', lp, fresh') -> + (fresh <= fresh')%positive. + Proof. + induction p; crush; repeat (destruct_match; []); simplify; + try (pose proof (IHp1 _ _ _ _ Heqp); pose proof (IHp2 _ _ _ _ Heqp4); lia). + apply IHp in Heqp0; lia. + Qed. + + Lemma unnest_predicate_lt_in_list : + forall p p' fresh fresh' lp, + unnest_predicate fresh p = (p', lp, fresh') -> + Forall (fun x => (fresh <= fst x < fresh')%positive) lp. + Proof. + induction p; crush; repeat (destruct_match; []); simplify. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; cbn. + + apply unnest_predicate_lt_fresh in Heqp4. + apply unnest_predicate_lt_fresh in Heqp. lia. + + apply Forall_app; split; apply Forall_forall; intros. + * eapply Forall_forall in H0; eauto. + apply unnest_predicate_lt_fresh in Heqp. lia. + * eapply Forall_forall in H; eauto. + apply unnest_predicate_lt_fresh in Heqp4. lia. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; cbn. + + apply unnest_predicate_lt_fresh in Heqp4. + apply unnest_predicate_lt_fresh in Heqp. lia. + + apply Forall_app; split; apply Forall_forall; intros. + * eapply Forall_forall in H0; eauto. + apply unnest_predicate_lt_fresh in Heqp. lia. + * eapply Forall_forall in H; eauto. + apply unnest_predicate_lt_fresh in Heqp4. lia. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; cbn. + + apply unnest_predicate_lt_fresh in Heqp4. + apply unnest_predicate_lt_fresh in Heqp. lia. + + apply Forall_app; split; apply Forall_forall; intros. + * eapply Forall_forall in H0; eauto. + apply unnest_predicate_lt_fresh in Heqp. lia. + * eapply Forall_forall in H; eauto. + apply unnest_predicate_lt_fresh in Heqp4. lia. + - pose proof (IHp _ _ _ _ Heqp0). + constructor; cbn. + + apply unnest_predicate_lt_fresh in Heqp0. lia. + + apply Forall_forall; intros. eapply Forall_forall in H; eauto. + lia. + Qed. + + Lemma unnest_predicate_lt_in_list5 : + forall p p' fresh fresh' lp, + unnest_predicate fresh p = (p', lp, fresh') -> + (max_hash_pred p < fresh)%positive -> + (max_hash_pred p' < fresh')%positive. + Proof. induction p; crush; repeat (destruct_match; []); crush. Qed. + + Lemma unnest_predicate_lt_in_list2 : + forall p p' fresh fresh' lp, + unnest_predicate fresh p = (p', lp, fresh') -> + (max_hash_pred p < fresh)%positive -> + Forall (fun x => (max_hash_pred (snd x) < fst x)%positive) lp. + Proof. + induction p; crush; repeat (destruct_match; []); simplify. + - pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). + pose proof (IHp1 _ _ _ _ Heqp ltac:(lia)) as X. + pose proof (IHp2 _ _ _ _ Heqp4 ltac:(lia)) as Y. + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp ltac:(lia)). + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp4 ltac:(lia)). + constructor; cbn; [lia|]. + apply Forall_app; split; eapply Forall_forall; intros. + + eapply Forall_forall in Y; eauto. + + eapply Forall_forall in X; eauto. + - pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). + pose proof (IHp1 _ _ _ _ Heqp ltac:(lia)) as X. + pose proof (IHp2 _ _ _ _ Heqp4 ltac:(lia)) as Y. + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp ltac:(lia)). + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp4 ltac:(lia)). + constructor; cbn; [lia|]. + apply Forall_app; split; eapply Forall_forall; intros. + + eapply Forall_forall in Y; eauto. + + eapply Forall_forall in X; eauto. + - pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). + pose proof (IHp1 _ _ _ _ Heqp ltac:(lia)) as X. + pose proof (IHp2 _ _ _ _ Heqp4 ltac:(lia)) as Y. + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp ltac:(lia)). + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp4 ltac:(lia)). + constructor; cbn; [lia|]. + apply Forall_app; split; eapply Forall_forall; intros. + + eapply Forall_forall in Y; eauto. + + eapply Forall_forall in X; eauto. + - pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp0). + pose proof (IHp _ _ _ _ Heqp0 ltac:(lia)) as X. + pose proof (unnest_predicate_lt_in_list5 _ _ _ _ _ Heqp0 ltac:(lia)). + constructor; cbn; [lia|auto]. + Qed. + + Lemma unnest_predicate_lt_in_list4 : + forall p p' fresh fresh' lp y, + unnest_predicate fresh p = (p', lp, fresh') -> + pred_In y p' -> + In y (map fst lp) \/ pred_In y p. + Proof. + induction p; crush; repeat (destruct_match; []); simplify. + - inv H0. tauto. + - inv H0. tauto. + - inv H0. tauto. + - inv H0. tauto. + Qed. + + Lemma unnest_predicate_lt_in_list3 : + forall p p' fresh fresh' lp, + unnest_predicate fresh p = (p', lp, fresh') -> + Forall (fun x => forall y, pred_In y (snd x) -> In y (map fst lp) \/ pred_In y p) lp. + Proof. + induction p; crush; repeat (destruct_match; []); simplify. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; intros; cbn [snd] in *. + + inv H1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Pand1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp4 H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Pand2. + + apply Forall_app; split. + * apply Forall_forall; intros. eapply Forall_forall in H0; eauto. + rewrite map_app. apply H0 in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Pand2. + * apply Forall_forall; intros. eapply Forall_forall in H; eauto. + rewrite map_app. apply H in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Pand1. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; intros; cbn [snd] in *. + + inv H1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Por1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp4 H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Por2. + + apply Forall_app; split. + * apply Forall_forall; intros. eapply Forall_forall in H0; eauto. + rewrite map_app. apply H0 in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Por2. + * apply Forall_forall; intros. eapply Forall_forall in H; eauto. + rewrite map_app. apply H in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Por1. + - pose proof (IHp1 _ _ _ _ Heqp). + pose proof (IHp2 _ _ _ _ Heqp4). + constructor; intros; cbn [snd] in *. + + inv H1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Pimp1. + * pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp4 H4). + inv H1. left. right. rewrite list_append_map. rewrite in_app. tauto. + right. now apply pred_In_Pimp2. + + apply Forall_app; split. + * apply Forall_forall; intros. eapply Forall_forall in H0; eauto. + rewrite map_app. apply H0 in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Pimp2. + * apply Forall_forall; intros. eapply Forall_forall in H; eauto. + rewrite map_app. apply H in H2. inv H2. + left. right. apply in_app; tauto. + right. now apply pred_In_Pimp1. + - pose proof (IHp _ _ _ _ Heqp0). + constructor; intros; cbn [snd] in *. + + inv H0. + pose proof (unnest_predicate_lt_in_list4 _ _ _ _ _ _ Heqp0 H3). + inv H0. tauto. + right. now apply pred_In_Pnot. + + apply Forall_forall; intros. eapply Forall_forall in H0; eauto. + apply H0 in H1. inv H1; [tauto|]. + right. now apply pred_In_Pnot. + Qed. + + Lemma unnest_predicate_sorted : + forall p p' lp fresh' fresh, + unnest_predicate fresh p = (p', lp, fresh') -> + Sorted.Sorted (Basics.flip pl_lt) lp. + Proof. + induction p; crush; try apply Sorted_nil; + repeat (destruct_match; []); simplify. + - constructor. + eapply SetoidList.SortA_app with (eqA := eq); auto. + eapply IHp2; eassumption. + eapply IHp1; eassumption. + intros. apply SetoidList.InA_alt in H; simplify. + apply SetoidList.InA_alt in H0; simplify. + unfold Basics.flip. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + eapply Forall_forall in H; [|eassumption]. + eapply Forall_forall in H0; [|eassumption]. + destruct x, x0; unfold pl_lt; cbn [snd fst] in *; lia. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + apply SetoidList.InfA_app. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H0; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. lia. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). lia. + - constructor. + eapply SetoidList.SortA_app with (eqA := eq); auto. + eapply IHp2; eassumption. + eapply IHp1; eassumption. + intros. apply SetoidList.InA_alt in H; simplify. + apply SetoidList.InA_alt in H0; simplify. + unfold Basics.flip. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + eapply Forall_forall in H; [|eassumption]. + eapply Forall_forall in H0; [|eassumption]. + destruct x, x0; unfold pl_lt; cbn [snd fst] in *; lia. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + apply SetoidList.InfA_app. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H0; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. lia. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). lia. + - constructor. + eapply SetoidList.SortA_app with (eqA := eq); auto. + eapply IHp2; eassumption. + eapply IHp1; eassumption. + intros. apply SetoidList.InA_alt in H; simplify. + apply SetoidList.InA_alt in H0; simplify. + unfold Basics.flip. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + eapply Forall_forall in H; [|eassumption]. + eapply Forall_forall in H0; [|eassumption]. + destruct x, x0; unfold pl_lt; cbn [snd fst] in *; lia. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp4). + apply SetoidList.InfA_app. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H0; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. lia. + + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H1; simplify. + eapply Forall_forall in H; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ Heqp4). lia. + - constructor. + eapply IHp; eassumption. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp0). + apply SetoidList.InA_InfA with (eqA := eq); auto; intros. + apply SetoidList.InA_alt in H0; simplify. + eapply Forall_forall in H; [|eassumption]. + unfold Basics.flip, pl_lt; destruct x; cbn [snd fst] in *. lia. + Qed. + + Lemma to_equiv_rev : + forall l, rev (to_equiv l) = to_equiv (rev l). + Proof. unfold to_equiv; intros. rewrite map_rev. auto. Qed. + + Lemma unnest_predicate_sorted_app : + forall p1 p l p0 p3 p4 l0 fresh p2, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + Sorted.Sorted (Basics.flip pl_lt) l -> + Sorted.Sorted (Basics.flip pl_lt) l0 -> + Sorted.Sorted (Basics.flip pl_lt) (l0 ++ l). + Proof. + intros * NEST1 NEST2 SORT1 SORT2. + apply SetoidList.SortA_app with (eqA := eq); auto. + intros. + apply SetoidList.InA_alt in H. + apply SetoidList.InA_alt in H0. simplify. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + eapply Forall_forall in H; [|eassumption]. + eapply Forall_forall in H1; [|eassumption]. + unfold Basics.flip, pl_lt. lia. + Qed. + + Lemma unnest_predicate_equiv_list : + forall p p0 p1 p2 p3 p4 l l0 fresh, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + (Pos.max (max_hash_pred p1) (max_hash_pred p2) < fresh)%positive -> + equiv_list_wf (rev (l0 ++ l)). + Proof. + intros * NEST1 NEST2 MAX. + unfold equiv_list_wf. split. unfold Sorted. + pose proof (unnest_predicate_sorted _ _ _ _ _ NEST1). + pose proof (unnest_predicate_sorted _ _ _ _ _ NEST2). + replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by (unfold Basics.flip; auto). + apply sorted_rev4. + eapply unnest_predicate_sorted_app; eauto. + apply Forall_rev; apply Forall_app; split; + eapply unnest_predicate_lt_in_list2; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + lia. lia. + Qed. + + Lemma pred_in_fold : + forall A l p y, + @pred_In A p (fold_left Pand l y) -> + (exists i, In i l /\ pred_In p i) \/ pred_In p y. + Proof. + induction l. + intros. simplify. tauto. + intros. simpl in *. apply IHl in H. + inv H; simplify. left. econstructor. + constructor. right. eauto. auto. + inv H0; try tauto. + do 2 econstructor; split; try eassumption; tauto. + Qed. + + Lemma max_hash_pred_lt : + forall p l, + pred_In l p -> + (l <= max_hash_pred p)%positive. + Proof. + induction p; intros; try inv H; crush; + try (apply IHp1 in H2; lia); + apply IHp2 in H2; lia. + Qed. + + Lemma unnest_predicate_not_in : + forall p p0 p1 p2 p3 p4 l l0 fresh y, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + (Pos.max (max_hash_pred p1) (max_hash_pred p2) < fresh)%positive -> + Forall (fun x : positive * hash_pred => + ~ pred_In (fst x) (fold_left Pand (to_equiv (l0 ++ l)) T)) + ((p4, y) :: nil). + Proof. + intros * NEST1 NEST2 MAX. + constructor; [|constructor]. cbn [fst]. unfold not; intros. + apply pred_in_fold in H. inv H; inv H0; + simplify. unfold to_equiv in H0. apply map_in_some in H0. simplify. + apply in_app_or in H0; inv H0. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TEMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + eapply Forall_forall in H0; [|eassumption]. + eapply Forall_forall in H2; [|eassumption]. destruct x0. cbn [snd fst] in *. + apply max_hash_pred_lt in H1. unfold equiv in H1. + unfold max_hash_pred in H1. fold max_hash_pred in H1. lia. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1). + eapply Forall_forall in H0; [|eassumption]. + eapply Forall_forall in H2; [|eassumption]. destruct x0. cbn [snd fst] in *. + apply max_hash_pred_lt in H1. unfold equiv in H1. + unfold max_hash_pred in H1. fold max_hash_pred in H1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + Qed. + + Lemma unnest_predicate_not_in4 : + forall p p0 p1 l fresh y, + unnest_predicate fresh p1 = (p, l, p0) -> + (max_hash_pred p1 < fresh)%positive -> + Forall (fun x : positive * hash_pred => + ~ pred_In (fst x) (fold_left Pand (to_equiv l) T)) + ((p0, y) :: nil). + Proof. + intros * NEST1 MAX. + constructor; [|constructor]. cbn [fst]. unfold not; intros. + apply pred_in_fold in H. inv H; inv H0; + simplify. unfold to_equiv in H0. apply map_in_some in H0. simplify. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1). + eapply Forall_forall in H; [|eassumption]. + eapply Forall_forall in H2; [|eassumption]. destruct x0. cbn [snd fst] in *. + apply max_hash_pred_lt in H1. unfold equiv in H1. + unfold max_hash_pred in H1. fold max_hash_pred in H1. lia. + Qed. + + Lemma unnest_predicate_in_pred : + forall fresh p1 p l p0, + unnest_predicate fresh p1 = (p, l, p0) -> + (max_hash_pred p1 < fresh)%positive -> + ~ pred_In p0 p. + Proof. + destruct p1; unfold not; intros. + - inv H. inv H1. + - inv H. inv H1. + - inv H. inv H1. + - inv H. crush. inv H1. lia. + - simplify. repeat (destruct_match; []). simplify. + inv H1. lia. + - simplify. repeat (destruct_match; []). simplify. + inv H1. lia. + - simplify. repeat (destruct_match; []). simplify. + inv H1. lia. + - simplify. repeat (destruct_match; []). simplify. + inv H1. lia. + Qed. + + Lemma unnest_predicate_in_pred2 : + forall fresh p1 p l p0, + unnest_predicate fresh p1 = (p, l, p0) -> + (max_hash_pred p1 < fresh)%positive -> + (max_hash_pred p < p0)%positive. + Proof. + intros. + destruct p1; crush. + - repeat (destruct_match; []); simplify; lia. + - repeat (destruct_match; []); simplify; lia. + - repeat (destruct_match; []); simplify; lia. + - repeat (destruct_match; []); simplify; lia. + Qed. + + Lemma unnest_predicate_forall_none : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∧ p2) < fresh)%positive -> + Forall (fun x : positive * predicate => a ! (fst x) = None) (rev (l0 ++ l)). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + apply Forall_forall; intros. apply GTMAX. + unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as X3. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as X4. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as X2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as Y1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y2. + apply in_rev in H. apply in_app in H. inv H. + eapply Forall_forall in X1; eauto; []. + eapply Forall_forall in X2; eauto; []. lia. + eapply Forall_forall in Y1; eauto; []. + eapply Forall_forall in Y2; eauto; []. lia. + Qed. + + Lemma unnest_predicate_forall_gt1 : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∧ p2) < fresh)%positive -> + Forall (fun x : positive * predicate => (1 < fst x)%positive) (rev (l0 ++ l)). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + apply Forall_forall; intros. + unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as X3. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as X4. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as X2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as Y1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y2. + apply in_rev in H. apply in_app in H. inv H. + eapply Forall_forall in X1; eauto; []. + eapply Forall_forall in X2; eauto; []. lia. + eapply Forall_forall in Y1; eauto; []. + eapply Forall_forall in Y2; eauto; []. lia. + Qed. + + Lemma unnest_predicate_not_in3 : + forall p1 p0 l fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred p1 < fresh)%positive -> + ~ In p0 (map fst (rev l)). + Proof. + intros * NEST1 EVAL1 GTMAX MAXHASH. + unfold not; intros. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + apply map_in_some in H. simplify. + eapply in_rev in H. + eapply Forall_forall in H0; eauto. + eapply Forall_forall in H1; eauto. + destruct x; cbn [fst snd max_hash_pred] in *. + lia. + Qed. + + Lemma unnest_predicate_not_in2 : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∧ p2) < fresh)%positive -> + ~ In p4 (map fst (rev l ++ rev l0)). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + unfold not; intros. + + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1). + apply map_in_some in H. simplify. + apply in_app_or in H. inv H. + eapply in_rev in H4. + eapply Forall_forall in H2; eauto. + eapply Forall_forall in H3; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + destruct x; cbn [fst snd] in *. lia. + eapply in_rev in H4. + eapply Forall_forall in H0; eauto. + eapply Forall_forall in H1; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + destruct x; cbn [fst snd] in *. lia. + Qed. + + Lemma eval_equiv_added : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∧ p2) < fresh)%positive -> + eval_hash_pred (equiv (Pbase p4) (p ∧ p3)) (upd_asgn (rev l ++ rev l0 ++ (p4, p ∧ p3) :: nil) a) = Some true. + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + apply eval_equiv2. rewrite eval_hash_pred_get. + rewrite app_assoc. + rewrite upd_asgn_get. + symmetry. rewrite upd_asgn_app1; auto. + constructor; [|constructor]; simpl. + pose proof NEST1 as X1. pose proof NEST2 as X2. + eapply unnest_predicate_in_pred2 in X1. + eapply unnest_predicate_in_pred2 in X2. + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + unfold not; intros. apply max_hash_pred_lt in H1. + cbn [max_hash_pred] in H1. lia. + + rewrite upd_asgn_not_in. eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + eapply unnest_predicate_not_in2; eauto. + Qed. + + Lemma eval_equiv_added_Por : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∨ p2) < fresh)%positive -> + eval_hash_pred (equiv (Pbase p4) (p ∨ p3)) (upd_asgn (rev l ++ rev l0 ++ (p4, p ∨ p3) :: nil) a) = Some true. + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + apply eval_equiv2. rewrite eval_hash_pred_get. + rewrite app_assoc. + rewrite upd_asgn_get. + symmetry. rewrite upd_asgn_app1; auto. + constructor; [|constructor]; simpl. + pose proof NEST1 as X1. pose proof NEST2 as X2. + eapply unnest_predicate_in_pred2 in X1. + eapply unnest_predicate_in_pred2 in X2. + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + unfold not; intros. apply max_hash_pred_lt in H1. + cbn [max_hash_pred] in H1. lia. + + rewrite upd_asgn_not_in. eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + eapply unnest_predicate_not_in2; eauto. + Qed. + + Lemma eval_equiv_added_Pimp : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 → p2) < fresh)%positive -> + eval_hash_pred (equiv (Pbase p4) (p → p3)) (upd_asgn (rev l ++ rev l0 ++ (p4, p → p3) :: nil) a) = Some true. + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + apply eval_equiv2. rewrite eval_hash_pred_get. + rewrite app_assoc. + rewrite upd_asgn_get. + symmetry. rewrite upd_asgn_app1; auto. + constructor; [|constructor]; simpl. + pose proof NEST1 as X1. pose proof NEST2 as X2. + eapply unnest_predicate_in_pred2 in X1. + eapply unnest_predicate_in_pred2 in X2. + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + 2: { unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + unfold not; intros. apply max_hash_pred_lt in H1. + cbn [max_hash_pred] in H1. lia. + + rewrite upd_asgn_not_in. eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + eapply unnest_predicate_not_in2; eauto. + Qed. + + Lemma unnest_predicate_correct_Pand : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∧ p2) < fresh)%positive -> + eval_hash_pred (p1 ∧ p2) a = + eval_hash_pred (Pbase p4 ∧ fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p ∧ p3))) + (upd_asgn (rev (l0 ++ l) ++ (p4, p ∧ p3) :: nil) a). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + assert (eval_hash_pred (fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p ∧ p3))) + (upd_asgn (rev l ++ rev l0 ++ (p4, p ∧ p3) :: nil) a) = Some true). + { apply eval_hash_pred_fold_Pand2 with (x := T). + rewrite app_assoc. + rewrite upd_asgn_app1; [|eapply unnest_predicate_not_in; eauto]. + replace (rev l ++ rev l0) with (rev (l0 ++ l)) by (rewrite rev_app_distr; auto). + apply fold_left_Pand_rev. + rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + eapply unnest_predicate_equiv_list; eauto. + eapply unnest_predicate_forall_none; eauto. + eapply unnest_predicate_forall_gt1; eauto. + rewrite eval_hash_pred_T_Pand. + eapply eval_equiv_added; eauto. + } + assert (eval_hash_pred (p1 ∧ p2) a = eval_hash_pred (Pbase p4) + (upd_asgn ((rev l ++ rev l0) ++ (p4, p ∧ p3) :: nil) a)). + { rewrite eval_hash_pred_get. + rewrite upd_asgn_get. + 2: { rewrite upd_asgn_not_in by (eapply unnest_predicate_not_in2; eauto). + eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + apply eval_hash_pred_pand. + { rewrite EVAL1. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app1. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + eapply Forall_forall in H1; eauto. eapply Forall_forall in H2; eauto. + unfold not; intros. apply max_hash_pred_lt in H5. destruct x; cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST1). + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST2). + lia. + } } + { rewrite EVAL2. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app2. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + apply Forall_rev. apply Forall_forall; intros. split. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + destruct x, x0. cbn [fst snd] in *. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [fst snd] in *. lia. + + pose proof (unnest_predicate_lt_in_list3 _ _ _ _ _ NEST2) as X. + unfold not; intros Y. eapply Forall_forall in X; eauto. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + destruct x, x0. cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + pose proof Y as YY. apply X in YY. inv YY. + eapply map_in_some in H4; simplify. destruct x; simplify. + eapply max_hash_pred_lt in Y. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2) as X22. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y22. + eapply Forall_forall in X22; eauto. eapply Forall_forall in Y22; eauto. + cbn [fst snd] in *. + cbn [max_hash_pred] in *. + lia. lia. + eapply max_hash_pred_lt in H4. lia. + } + { intros. apply in_rev in H0. destruct x; cbn [fst snd] in *. + unfold not; intros Y. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [max_hash_pred fst snd] in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + exploit unnest_predicate_lt_in_list4. eapply NEST2. eauto. intros. inv H3. + eapply map_in_some in H4; simplify. + destruct x. + cbn [max_hash_pred fst snd] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + cbn [max_hash_pred fst snd] in *. + eapply max_hash_pred_lt in Y. lia. + eapply max_hash_pred_lt in H4. lia. + } } } + symmetry; rewrite eval_hash_pred_T_Pand2; rewrite rev_app_distr; + try solve [rewrite <- app_assoc; auto]; auto. + Qed. + + Lemma unnest_predicate_correct_Por : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 ∨ p2) < fresh)%positive -> + eval_hash_pred (p1 ∨ p2) a = + eval_hash_pred (Pbase p4 ∧ fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p ∨ p3))) + (upd_asgn (rev (l0 ++ l) ++ (p4, p ∨ p3) :: nil) a). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + assert (eval_hash_pred (fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p ∨ p3))) + (upd_asgn (rev l ++ rev l0 ++ (p4, p ∨ p3) :: nil) a) = Some true). + { apply eval_hash_pred_fold_Pand2 with (x := T). + rewrite app_assoc. + rewrite upd_asgn_app1; [|eapply unnest_predicate_not_in; eauto]. + replace (rev l ++ rev l0) with (rev (l0 ++ l)) by (rewrite rev_app_distr; auto). + apply fold_left_Pand_rev. + rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + eapply unnest_predicate_equiv_list; eauto. + eapply unnest_predicate_forall_none; eauto. + eapply unnest_predicate_forall_gt1; eauto. + rewrite eval_hash_pred_T_Pand. + eapply eval_equiv_added_Por; eauto. + } + assert (eval_hash_pred (p1 ∨ p2) a = eval_hash_pred (Pbase p4) + (upd_asgn ((rev l ++ rev l0) ++ (p4, p ∨ p3) :: nil) a)). + { rewrite eval_hash_pred_get. + rewrite upd_asgn_get. + 2: { rewrite upd_asgn_not_in by (eapply unnest_predicate_not_in2; eauto). + eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + apply eval_hash_pred_por. + { rewrite EVAL1. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app1. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + eapply Forall_forall in H1; eauto. eapply Forall_forall in H2; eauto. + unfold not; intros. apply max_hash_pred_lt in H5. destruct x; cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST1). + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST2). + lia. + } } + { rewrite EVAL2. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app2. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + apply Forall_rev. apply Forall_forall; intros. split. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + destruct x, x0. cbn [fst snd] in *. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [fst snd] in *. lia. + + pose proof (unnest_predicate_lt_in_list3 _ _ _ _ _ NEST2) as X. + unfold not; intros Y. eapply Forall_forall in X; eauto. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + destruct x, x0. cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + pose proof Y as YY. apply X in YY. inv YY. + eapply map_in_some in H4; simplify. destruct x; simplify. + eapply max_hash_pred_lt in Y. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2) as X22. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y22. + eapply Forall_forall in X22; eauto. eapply Forall_forall in Y22; eauto. + cbn [fst snd] in *. + cbn [max_hash_pred] in *. + lia. lia. + eapply max_hash_pred_lt in H4. lia. + } + { intros. apply in_rev in H0. destruct x; cbn [fst snd] in *. + unfold not; intros Y. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [max_hash_pred fst snd] in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + exploit unnest_predicate_lt_in_list4. eapply NEST2. eauto. intros. inv H3. + eapply map_in_some in H4; simplify. + destruct x. + cbn [max_hash_pred fst snd] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + cbn [max_hash_pred fst snd] in *. + eapply max_hash_pred_lt in Y. lia. + eapply max_hash_pred_lt in H4. lia. + } } } + symmetry; rewrite eval_hash_pred_T_Pand2; rewrite rev_app_distr; + try solve [rewrite <- app_assoc; auto]; auto. + Qed. + +Lemma unnest_predicate_correct_Pimp : + forall p1 p2 p0 p3 l0 l p4 fresh p a, + unnest_predicate fresh p1 = (p, l, p0) -> + unnest_predicate p0 p2 = (p3, l0, p4) -> + eval_hash_pred p1 a = eval_pred_list (p, l) (upd_asgn (rev l) a) -> + eval_hash_pred p2 a = eval_pred_list (p3, l0) (upd_asgn (rev l0) a) -> + (forall y, (fresh <= y)%positive -> a ! y = None) -> + (max_hash_pred (p1 → p2) < fresh)%positive -> + eval_hash_pred (p1 → p2) a = + eval_hash_pred (Pbase p4 ∧ fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p → p3))) + (upd_asgn (rev (l0 ++ l) ++ (p4, p → p3) :: nil) a). + Proof. + intros * NEST1 NEST2 EVAL1 EVAL2 GTMAX MAXHASH. + assert (eval_hash_pred (fold_left Pand (to_equiv (l0 ++ l)) (T ∧ equiv (Pbase p4) (p → p3))) + (upd_asgn (rev l ++ rev l0 ++ (p4, p → p3) :: nil) a) = Some true). + { apply eval_hash_pred_fold_Pand2 with (x := T). + rewrite app_assoc. + rewrite upd_asgn_app1; [|eapply unnest_predicate_not_in; eauto]. + replace (rev l ++ rev l0) with (rev (l0 ++ l)) by (rewrite rev_app_distr; auto). + apply fold_left_Pand_rev. + rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + eapply unnest_predicate_equiv_list; eauto. + eapply unnest_predicate_forall_none; eauto. + eapply unnest_predicate_forall_gt1; eauto. + rewrite eval_hash_pred_T_Pand. + eapply eval_equiv_added_Pimp; eauto. + } + assert (eval_hash_pred (p1 → p2) a = eval_hash_pred (Pbase p4) + (upd_asgn ((rev l ++ rev l0) ++ (p4, p → p3) :: nil) a)). + { rewrite eval_hash_pred_get. + rewrite upd_asgn_get. + 2: { rewrite upd_asgn_not_in by (eapply unnest_predicate_not_in2; eauto). + eapply GTMAX. unfold max_hash_pred in *. fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). lia. + } + apply eval_hash_pred_pimp. + { rewrite EVAL1. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app1. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. + cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)). + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2). + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + eapply Forall_forall in H1; eauto. eapply Forall_forall in H2; eauto. + unfold not; intros. apply max_hash_pred_lt in H5. destruct x; cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST1). + pose proof (unnest_predicate_in_pred2 _ _ _ _ _ NEST2). + lia. + } } + { rewrite EVAL2. unfold eval_pred_list. cbn [snd fst]. + rewrite upd_asgn_app2. + { rewrite eval_hash_pred_T_Pand2; auto. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + { unfold equiv_list_wf, Sorted; split. + { replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. } + { apply Forall_rev. eapply unnest_predicate_lt_in_list2; eauto. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + cbn [max_hash_pred] in *; lia. } } + { exploit unnest_predicate_forall_none. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } + { exploit unnest_predicate_forall_gt1. apply NEST1. apply NEST2. + all: eauto. intros. cbn. rewrite rev_app_distr in H0. + apply Forall_app in H0; crush. } } + { apply Forall_rev. apply Forall_forall; intros. + apply Forall_rev. apply Forall_forall; intros. split. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + destruct x, x0. cbn [fst snd] in *. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [fst snd] in *. lia. + + pose proof (unnest_predicate_lt_in_list3 _ _ _ _ _ NEST2) as X. + unfold not; intros Y. eapply Forall_forall in X; eauto. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as TMP. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + destruct x, x0. cbn [fst snd] in *. + unfold max_hash_pred in *; fold max_hash_pred in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + pose proof Y as YY. apply X in YY. inv YY. + eapply map_in_some in H4; simplify. destruct x; simplify. + eapply max_hash_pred_lt in Y. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2) as X22. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y22. + eapply Forall_forall in X22; eauto. eapply Forall_forall in Y22; eauto. + cbn [fst snd] in *. + cbn [max_hash_pred] in *. + lia. lia. + eapply max_hash_pred_lt in H4. lia. + } + { intros. apply in_rev in H0. destruct x; cbn [fst snd] in *. + unfold not; intros Y. + cbn [max_hash_pred] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST1 ltac:(lia)) as X1. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST1) as Y1. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1) as Z1. + eapply Forall_forall in X1; eauto. eapply Forall_forall in Y1; eauto. + cbn [max_hash_pred fst snd] in *. + exploit unnest_predicate_in_pred2; eauto. lia; intros. intros. + exploit unnest_predicate_in_pred2. eapply NEST1. lia. intros. + exploit unnest_predicate_lt_in_list4. eapply NEST2. eauto. intros. inv H3. + eapply map_in_some in H4; simplify. + destruct x. + cbn [max_hash_pred fst snd] in *. + pose proof (unnest_predicate_lt_in_list2 _ _ _ _ _ NEST2 ltac:(lia)) as X2. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ NEST2) as Y2. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST2) as Z2. + eapply Forall_forall in X2; eauto. eapply Forall_forall in Y2; eauto. + cbn [max_hash_pred fst snd] in *. + eapply max_hash_pred_lt in Y. lia. + eapply max_hash_pred_lt in H4. lia. + } } } + symmetry; rewrite eval_hash_pred_T_Pand2; rewrite rev_app_distr; + try solve [rewrite <- app_assoc; auto]; auto. + Qed. + + Lemma upd_asgn_nil : + forall a, upd_asgn nil a = a. + Proof. crush. Qed. + + #[local] Opaque eval_hash_pred. + #[local] Opaque upd_asgn. + + Lemma unnest_predicate_correct : + forall a p fresh p' lp fresh', + (max_hash_pred p < fresh)%positive -> + unnest_predicate fresh p = (p', lp, fresh') -> + (forall y : positive, (fresh <= y)%positive -> a ! y = None) -> + eval_hash_pred p a = eval_pred_list (p', lp) (upd_asgn (rev lp) a). + Proof. + induction p; try solve [crush]. + - unfold eval_pred_list. simplify. + rewrite eval_hash_pred_T_Pand3. + rewrite upd_asgn_nil. auto. + - simplify. + destruct (unnest_predicate fresh p1) eqn:?. destruct p. + destruct (unnest_predicate p0 p2) eqn:?. destruct p3. + unfold eval_pred_list. simplify. + pose proof Heqp as X1. pose proof Heqp3 as X2. + pose proof X2 as X3. + pose proof X1 as X4. + apply IHp1 in X1; [|lia|auto]. + apply unnest_predicate_lt_fresh in X2. + apply unnest_predicate_lt_fresh in X4. + apply IHp2 in X3; [|lia|intros; eapply H1; lia]. + eapply unnest_predicate_correct_Pand; eauto. + - simplify. + destruct (unnest_predicate fresh p1) eqn:?. destruct p. + destruct (unnest_predicate p0 p2) eqn:?. destruct p3. + unfold eval_pred_list. simplify. + pose proof Heqp as X1. pose proof Heqp3 as X2. + pose proof X2 as X3. + pose proof X1 as X4. + apply IHp1 in X1; [|lia|auto]. + apply unnest_predicate_lt_fresh in X2. + apply unnest_predicate_lt_fresh in X4. + apply IHp2 in X3; [|lia|intros; eapply H1; lia]. + eapply unnest_predicate_correct_Por; eauto. + - simplify. + destruct (unnest_predicate fresh p1) eqn:?. destruct p. + destruct (unnest_predicate p0 p2) eqn:?. destruct p3. + unfold eval_pred_list. simplify. + pose proof Heqp as X1. pose proof Heqp3 as X2. + pose proof X2 as X3. + pose proof X1 as X4. + apply IHp1 in X1; [|lia|auto]. + apply unnest_predicate_lt_fresh in X2. + apply unnest_predicate_lt_fresh in X4. + apply IHp2 in X3; [|lia|intros; eapply H1; lia]. + eapply unnest_predicate_correct_Pimp; eauto. + - simplify. repeat destruct_match; simplify. + unfold eval_pred_list; simplify. + pose proof Heqp0 as X1. + pose proof Heqp0 as X2. + apply IHp in X1; [|lia|auto]. + apply unnest_predicate_lt_fresh in X2. + symmetry; rewrite eval_hash_pred_T_Pand2; [symmetry|]. + + rewrite eval_hash_pred_get. + rewrite upd_asgn_get; + [|rewrite upd_asgn_not_in; auto; + eapply unnest_predicate_not_in3; eauto]; []. + apply eval_hash_pred_pnot. rewrite X1. unfold eval_pred_list. + simplify. + rewrite eval_hash_pred_T_Pand2; auto; []. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + * unfold equiv_list_wf; split. + ** unfold Sorted. + replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. + ** apply Forall_rev; eapply unnest_predicate_lt_in_list2; eauto. + * apply Forall_rev. apply Forall_forall; intros. apply H1. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + lia. + * apply Forall_rev. apply Forall_forall; intros. simplify. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + lia. + + apply eval_hash_pred_fold_Pand2 with (x := T). + * rewrite upd_asgn_app1; [|eapply unnest_predicate_not_in4; eauto]; []. + apply fold_left_Pand_rev. rewrite to_equiv_rev. + apply eval_hash_pred_equiv_asgn; auto. + ** unfold equiv_list_wf; split. + *** unfold Sorted. + replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + apply sorted_rev4. eapply unnest_predicate_sorted; eauto. + *** apply Forall_rev; eapply unnest_predicate_lt_in_list2; eauto. + ** apply Forall_rev. apply Forall_forall; intros. apply H1. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + lia. + ** apply Forall_rev. apply Forall_forall; intros. simplify. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + lia. + * rewrite eval_hash_pred_T_Pand. + apply eval_equiv2. rewrite eval_hash_pred_get. + rewrite upd_asgn_get. + ** symmetry. rewrite upd_asgn_app1; auto. + constructor; [|constructor]; simpl. + pose proof Heqp0 as NEST1. + pose proof NEST1 as Y1. + eapply unnest_predicate_in_pred2 in Y1. + *** pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). + unfold not; intros. apply max_hash_pred_lt in H2. + cbn [max_hash_pred] in H2. lia. + *** unfold max_hash_pred in *; fold max_hash_pred in *. + pose proof (unnest_predicate_lt_fresh _ _ _ _ _ NEST1). lia. + ** rewrite upd_asgn_not_in; auto. eapply unnest_predicate_not_in3; eauto. + Qed. + +#[local] Transparent eval_hash_pred. +#[local] Transparent upd_asgn. + +Lemma interp_aux_correct : + forall a ba f h1 h2, + (forall l, h1 <> Fimp l) -> + is_true (form_eqb h1 h2) -> + interp_aux a ba f h1 = interp_aux a ba f h2. +Proof. + unfold is_true. intros * IMP **. + destruct h1; unfold form_eqb in *; try destruct_match; crush; auto. + - apply Int63.eqb_spec in H; subst; auto. + - apply Int63.eqb_spec in H0; subst; + apply Int63.eqb_spec in H1; subst; auto. + - apply list_beq_spec in H. subst. + rewrite !Misc.afold_left_and. rewrite H. auto. + exact Int63.eqb_spec. + - apply list_beq_spec in H. subst. + rewrite !Misc.afold_left_or. rewrite H. auto. + exact Int63.eqb_spec. + - apply Int63.eqb_spec in H0. apply Int63.eqb_spec in H1. subst; auto. + - apply Int63.eqb_spec in H0. apply Int63.eqb_spec in H1. subst; auto. + - apply Int63.eqb_spec in H. apply Int63.eqb_spec in H1. apply Int63.eqb_spec in H2. subst; auto. +Qed. + +Lemma form_eqb_symm : + forall f1 f2, + is_true (form_eqb f1 f2) -> + is_true (form_eqb f2 f1). +Proof. + unfold is_true, form_eqb. intros. repeat destruct_match; try discriminate; auto; + repeat match goal with + | H: Int63.eqb _ _ = true |- _ => apply Int63.eqb_spec in H + | |- Int63.eqb _ _ = true => apply Int63.eqb_spec + | H: list_beq _ _ _ = true |- _ => apply list_beq_spec in H; [|exact Int63.eqb_spec] + | |- list_beq _ _ _ = true => apply list_beq_spec; [exact Int63.eqb_spec|] + | H: _ && _ = true |- _ => simplify + | |- _ && _ = true => apply andb_true_intro; split + end; subst; auto. +Qed. + +Lemma t_interp_check_form : (* First array might be shorter, but it's ok *) + forall a ba f1 f2 l, + is_true (check_array form_eqb f1 f2) -> + is_true (check_local_form l f1) -> + forall i, + to_Z i < to_Z (PArray.length f1) -> + PArray.get (t_interp a ba f1) i = PArray.get (t_interp a ba f2) i. +Proof. + intros * TRUE CHK. + pose proof (check_array_length _ _ _ _ TRUE) as BOUNDS. + pose proof (length_t_interp a ba f1) as Y. + pose proof (length_t_interp a ba f2) as Z. + revert Y. + unfold t_interp. + apply Misc.foldi_ind with (P := + fun iter fol => + PArray.length fol = PArray.length f1 -> + forall i, + to_Z i < to_Z iter -> + PArray.get fol i = PArray.get _ i + ); intros. + + apply Misc.leb_0. + + pose proof (not_int63_lt_0 i). replace (to_Z 0) with 0 in H0 by auto. lia. + + rewrite PArray.length_set in H2. specialize (H1 H2). + erewrite Misc.to_Z_add_1 in H3 by eassumption. + destruct (int_dec i i0); subst. + - rewrite PArray.get_set_same by (now rewrite H2). + revert H1. revert Z. unfold t_interp. + assert (W: to_Z i0 < to_Z (PArray.length f2)). + { apply Int63.ltb_spec in H0. lia. } revert W. + apply Misc.foldi_ind with (P := + fun iter fol => + to_Z i0 < to_Z iter -> + PArray.length fol = PArray.length f2 -> + (forall i, to_Z i < to_Z i0 -> + PArray.get a0 i = PArray.get fol i) -> + _ = PArray.get fol i0 + ); intros. + * apply Misc.leb_0. + * pose proof (not_int63_lt_0 i0). replace (to_Z 0) with 0 in H1 by auto. lia. + * rewrite PArray.length_set in H7. erewrite Misc.to_Z_add_1 in H6 by eassumption. + destruct (int_dec i0 i); subst. rewrite PArray.get_set_same by (now rewrite H7). + rewrite lt_form_interp_form_aux with (f2 := PArray.get a1) (i := i). + erewrite interp_aux_correct; auto. + eapply check_local_form_spec; eauto. + apply form_eqb_symm. eapply check_array_correct; eauto. + apply Int63.ltb_spec in H0. auto. + intros. rewrite H8. rewrite PArray.get_set_other; auto. + unfold not; intros; subst. apply Int63.ltb_spec in H9. lia. + apply Int63.ltb_spec in H9; auto. + unfold check_local_form, is_true in *. simplify. + unfold check_form in *. simplify. unfold wf in *. + eapply Misc.aforallbi_spec in H12; eauto. + rewrite PArray.get_set_other by auto. eapply H5; auto. + assert (to_Z i <> to_Z i0) by (unfold not; intros; apply n; now apply to_Z_inj). + lia. intros. rewrite H8; auto. rewrite PArray.get_set_other. auto. + unfold not; intros. subst. lia. + - rewrite PArray.get_set_other by auto. eapply H1. + assert (to_Z i <> to_Z i0) by (unfold not; intros; apply n; now apply to_Z_inj). + lia. +Qed. + +(* Lemma a_interp_check_form : *) +(* forall t_i (t_func: PArray.array (tval t_i)) a ba a1 a2, *) +(* is_true (check_array Atom.eqb a1 a2) -> *) +(* is_true (check_array Bool.eqb (Atom.t_interp t_i t_func a1) (Atom.t_interp t_i t_func a2)). *) + +(* Misc.foldi_ind2: *) +(* forall (A B : Type) (P : Int63.int -> A -> B -> Prop) (f1 : Int63.int -> A -> A) *) +(* (f2 : Int63.int -> B -> B) (from to : Int63.int) (a1 : A) (a2 : B), *) +(* Int63.leb from to = true -> *) +(* P from a1 a2 -> *) +(* (forall (i : Int63.int) (a3 : A) (a4 : B), *) +(* Int63.leb from i = true -> Int63.ltb i to = true -> P i a3 a4 -> P (Int63.add i 1) (f1 i a3) (f2 i a4)) -> *) +(* P to (Misc.foldi f1 from to a1) (Misc.foldi f2 from to a2) *) + +Definition ind_S A a_init j (to: Int63.int) (fold1: PArray.array A) := + PArray.get fold1 j = PArray.get a_init j. + +Lemma ind_S_nochange : + forall a2 j from to a_init, + Int63.to_Z j < Int63.to_Z from <= Int63.to_Z to -> + ind_S _ a_init j to (Misc.foldi + (fun (i : Int63.int) (t_a : PArray.array (bval t_i)) => + PArray.set t_a i (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a2 i))) from to a_init). +Proof. + intros. apply Misc.foldi_ind. apply Int63.leb_spec. lia. now unfold ind_S. + intros. unfold ind_S. rewrite PArray.get_set_other. + unfold ind_S in H2. auto. apply Int63.leb_spec in H0. + unfold not. intros. assert (Int63.to_Z i = Int63.to_Z j) by (now subst). + lia. +Qed. + +Lemma foldi_array_get_constant : + forall a2 j from to a_init, + Int63.to_Z j < Int63.to_Z from <= Int63.to_Z to -> + PArray.get (Misc.foldi + (fun (i : Int63.int) (t_a : PArray.array (bval t_i)) => + PArray.set t_a i (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a2 i))) from to a_init) j = PArray.get a_init j. +Proof. apply ind_S_nochange; auto. Qed. + +Definition interp_ind_S t_i j (to: Int63.int) (fold1: PArray.array (bval t_i)) fold2 := + PArray.get fold1 j = PArray.get fold2 j. + +(* (Misc.foldi *) +(* (fun (i : Int63.int) (t_a : PArray.array (bval t_i)) => *) +(* PArray.set t_a i (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a2 i))) (Int63.of_Z 0) *) +(* (PArray.length a2) (PArray.make (PArray.length a2) (interp_cop t_i CO_xH))) *) + +Lemma get_t_interp : + forall a1 a2 j, + is_true (Atom.wf a1) -> + is_true (check_array Atom.eqb a1 a2) -> + Int63.to_Z j < Int63.to_Z (PArray.length a1) -> + PArray.get (Atom.t_interp t_i t_func a1) j = + PArray.get (Atom.t_interp t_i t_func a2) j. +Proof. + intros * WF CHECK BOUNDS. unfold Atom.t_interp. + pose proof (check_array_length _ _ _ _ CHECK) as ABOUNDS. + pose proof (check_array_correct _ _ _ _ CHECK). + generalize dependent j. + assert (TMP: forall (A: Prop), + A /\ ((Int63.to_Z (PArray.length (Misc.foldi + (fun (i : Int63.int) (t_a : PArray.array (bval t_i)) => + PArray.set t_a i (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a1 i))) (Int63.of_Z 0) + (PArray.length a1) (PArray.make (PArray.length a1) (interp_cop t_i CO_xH))))) + = Int63.to_Z (PArray.length a1)) -> A) by tauto. apply TMP. clear TMP. + apply Misc.foldi_ind with + (P := fun iter fol => + (forall j, + Int63.to_Z j < Int63.to_Z iter -> + PArray.get fol j = + PArray.get + (Misc.foldi + (fun (i : Int63.int) (t_a : PArray.array (bval t_i)) => + PArray.set t_a i (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a2 i))) + (Int63.of_Z 0) + (PArray.length a2) (PArray.make (PArray.length a2) (interp_cop t_i CO_xH))) j) /\ + Int63.to_Z (PArray.length fol) = Int63.to_Z (PArray.length a1) + ). + + apply Int63.leb_spec; simpl. + pose proof (not_int63_lt_0 (PArray.length a1)). + rewrite Int63.to_Z_0. auto. + + split. + + intros j BOUNDS0. pose proof (not_int63_lt_0 j). simpl in BOUNDS0. rewrite Int63.to_Z_0 in BOUNDS0. lia. + auto. + + intros i a ILEB ILTB [HIND ARRLEN]; split. + + intros j BOUNDS1. + erewrite Misc.to_Z_add_1 in BOUNDS1 by eassumption. + apply Int63.ltb_spec in ILTB. + apply Int63.leb_spec in ILEB. + assert (X: Int63.to_Z j < Int63.to_Z i \/ Int63.to_Z j = Int63.to_Z i) by lia. + inversion X as [X1|X1]. rewrite PArray.get_set_other. + rewrite HIND; auto; try lia. + unfold not; intro EQ. assert (Int63.to_Z j = Int63.to_Z i) by (now subst). lia. + assert (j = i) by (now apply Int63.to_Z_inj). subst. + rewrite PArray.get_set_same. 2: { apply Int63.ltb_spec. lia. } + generalize dependent HIND. + assert (X2: Int63.to_Z i < Int63.to_Z (PArray.length a2)) by lia. + generalize X2. + assert (TMP: forall (A: Prop), + A /\ ((Int63.to_Z (PArray.length (Misc.foldi + (fun (i0 : Int63.int) (t_a : PArray.array (bval t_i)) => + PArray.set t_a i0 (Atom.interp_aux t_i t_func (PArray.get t_a) (PArray.get a2 i0))) + (Int63.of_Z 0) (PArray.length a2) (PArray.make (PArray.length a2) (interp_cop t_i CO_xH))))) + = Int63.to_Z (PArray.length a2)) -> A) by tauto. apply TMP. clear TMP. + apply Misc.foldi_ind with + (P := fun iter fol => + (Int63.to_Z i < Int63.to_Z iter -> + (forall j : Int63.int, + Int63.to_Z j < Int63.to_Z i -> + PArray.get a j = PArray.get fol j) -> + Atom.interp_aux t_i t_func (PArray.get a) (PArray.get a1 i) = PArray.get fol i) + /\ Int63.to_Z (PArray.length fol) = Int63.to_Z (PArray.length a2) + ). + + apply Int63.leb_spec. lia. + + split. intros BOUNDS2. + pose proof (not_int63_lt_0 i). lia. auto. + + intros i0 a0 ILEB1 ILTB1 HIND2. split. + intros BOUNDS3 HHIND. + erewrite Misc.to_Z_add_1 in BOUNDS3 by eassumption. + assert (A: Int63.to_Z i < Int63.to_Z i0 \/ Int63.to_Z i = Int63.to_Z i0) by lia. + inversion A as [A1|A1]; clear A; subst. + rewrite PArray.get_set_other. eapply HIND2. lia. + intros j BOUNDS4. rewrite HHIND. rewrite PArray.get_set_other; auto. + unfold not; intros EQ. assert (U: Int63.to_Z j = Int63.to_Z i0) by (now subst). lia. lia. + unfold not; intros EQ. assert (U: Int63.to_Z i = Int63.to_Z i0) by (now subst). lia. + assert (U: i = i0) by (now apply Int63.to_Z_inj). subst. + rewrite PArray.get_set_same by (apply Int63.ltb_spec; lia). + assert (EQ: PArray.get a2 i0 = PArray.get a1 i0). + { apply Atom.eqb_spec. apply check_array_correct; auto. } + + rewrite EQ in *. + apply lt_interp_aux with (i := i0). intros j ILTB2. + rewrite HHIND. apply Int63.ltb_spec in ILTB2. rewrite PArray.get_set_other. + auto. unfold not; intros EQ1. assert (U: Int63.to_Z i0 = Int63.to_Z j) by (now subst). lia. + apply Int63.ltb_spec in ILTB2. + auto. + unfold Atom.wf in WF. eapply Misc.aforallbi_spec in WF. eassumption. + apply Int63.ltb_spec. lia. + rewrite PArray.length_set. lia. + rewrite PArray.length_set. lia. +Qed. + +Lemma interp_form_hatom_preserved : + forall a1 a2 x, + is_true (Atom.wf a1) -> + is_true (check_array Atom.eqb a1 a2) -> + Int63.to_Z x < Int63.to_Z (PArray.length a1) -> + interp_form_hatom t_i t_func a1 x = interp_form_hatom t_i t_func a2 x. +Proof. + unfold interp_form_hatom. unfold interp_hatom. intros. + erewrite get_t_interp; auto. +Qed. + +(* Definition F f2 := *) +(* (fun i fol => *) +(* (forall x, 0 <= Int63.to_Z x < Int63.to_Z i -> *) +(* PArray.get fol x *) +(* = PArray.get *) +(* (Misc.foldi *) +(* (fun (i : Int63.int) (t_b : PArray.array bool) => *) +(* PArray.set t_b i *) +(* (interp_aux (interp_form_hatom t_i t_func a1) *) +(* (interp_form_hatom_bv t_i t_func a1) *) +(* (PArray.get t_b) (PArray.get f2 i))) *) +(* (Int63.of_Z 0) (PArray.length f2) *) +(* (PArray.make (PArray.length f2) true)) x)). *) + +Definition form_atom f P := + forall i a, PArray.get f i = Fatom a -> P a. + +Definition single_form_atom f P := + forall a, f = Fatom a -> P a. + +Lemma lt_form_interp_single_form_atom : + forall a1 a2 b f h, + single_form_atom h (fun j => a1 j = a2 j) -> + interp_aux a1 b f h = interp_aux a2 b f h. +Proof. + destruct h;simpl;intros; trivial. auto. +Qed. + +Lemma lt_form_interp_not_bv : + forall b1 b2 a f h, + is_true (single_form_not_bv h) -> + interp_aux a b1 f h = interp_aux a b2 f h. +Proof. + destruct h;simpl;intros; trivial. discriminate. +Qed. + +Lemma check_single_form_not_bv : + forall f1 l x, + is_true (check_local_form l f1) -> + to_Z x < to_Z (PArray.length f1) -> + is_true (single_form_not_bv (PArray.get f1 x)). +Proof. + unfold check_local_form, is_true; simplify. + eapply Misc.aforallbi_spec in H2; [|apply Int63.ltb_spec; eassumption]. + crush. +Qed. + +Lemma is_true_lt_form : + forall a b f1 x, + to_Z a < to_Z b -> + is_true (lt_form a (PArray.get f1 x)) -> + is_true (lt_form b (PArray.get f1 x)). +Proof. + unfold is_true, lt_form; intros. destruct_match; auto. + - apply Int63.ltb_spec in H0. apply Int63.ltb_spec. lia. + - apply Misc.aforallbi_spec; intros. eapply Misc.aforallbi_spec in H0; eauto. + apply Int63.ltb_spec in H0. apply Int63.ltb_spec. lia. + - apply Misc.aforallbi_spec; intros. eapply Misc.aforallbi_spec in H0; eauto. + apply Int63.ltb_spec in H0. apply Int63.ltb_spec. lia. + - apply Misc.aforallbi_spec; intros. eapply Misc.aforallbi_spec in H0; eauto. + apply Int63.ltb_spec in H0. apply Int63.ltb_spec. lia. + - simplify. apply Int63.ltb_spec in H1. apply Int63.ltb_spec in H2. + apply andb_true_intro. split; apply Int63.ltb_spec; lia. + - simplify. apply Int63.ltb_spec in H1. apply Int63.ltb_spec in H2. + apply andb_true_intro. split; apply Int63.ltb_spec; lia. + - simplify. apply Int63.ltb_spec in H2. apply Int63.ltb_spec in H3. + apply Int63.ltb_spec in H0. + apply andb_true_intro; split; [apply andb_true_intro;split|]; + apply Int63.ltb_spec; lia. + - apply forallb_forall; intros. eapply forallb_forall in H0; try eassumption. + apply Int63.ltb_spec in H0. apply Int63.ltb_spec. lia. +Qed. + +Lemma get_t_interp_form : + forall a1 a2 f1, + is_true (check_array Atom.eqb a1 a2) -> + is_true (check_local_form (PArray.length a1) f1) -> + is_true (Atom.wf a1) -> + forall x, + Int63.to_Z x < Int63.to_Z (PArray.length f1) -> + (PArray.get (t_interp (interp_form_hatom t_i t_func a2) (interp_form_hatom_bv t_i t_func a2) f1) x + = PArray.get (t_interp (interp_form_hatom t_i t_func a1) (interp_form_hatom_bv t_i t_func a1) f1) x). +Proof. + intros * TRUE LOCAL WFA. + pose proof (length_t_interp (interp_form_hatom t_i t_func a1) + (interp_form_hatom_bv t_i t_func a1) f1) as Y1. + pose proof (length_t_interp (interp_form_hatom t_i t_func a2) + (interp_form_hatom_bv t_i t_func a2) f1) as Y2. + revert Y1. revert Y2. + unfold t_interp. + apply Misc.foldi_ind2 with (P := + fun i a b => + PArray.length a = PArray.length f1 -> + PArray.length b = PArray.length f1 -> + forall x, + to_Z x < to_Z (PArray.length f1) -> + PArray.get a x = PArray.get b x + ). + - apply Misc.leb_0. + - intros. now rewrite PArray.get_make. + - intros. rewrite PArray.length_set in H2. rewrite PArray.length_set in H3. + specialize (H1 H2 H3). + destruct (int_dec i x); subst. + + symmetry. rewrite !PArray.get_set_same + by (try rewrite H3 in *; try rewrite H2 in *; apply Int63.ltb_spec; lia); auto. + erewrite lt_form_interp_not_bv; [|eapply check_single_form_not_bv; eauto]. + erewrite lt_form_interp_single_form_atom. + 2: { unfold single_form_atom; intros. + rewrite interp_form_hatom_preserved with (a2 := a2). eauto. + auto. auto. apply check_local_form_spec in LOCAL; simplify. + eapply Int63.ltb_spec. eapply H7; eauto. } + erewrite lt_form_interp_form_aux. eauto. intros. symmetry. eapply H1; eauto. + apply Int63.ltb_spec in H5. eauto. + unfold is_true in LOCAL. unfold check_local_form in LOCAL. simplify. unfold check_form in H5. + simplify. unfold wf in H8. eapply Misc.aforallbi_spec in H8; eauto. + eapply is_true_lt_form; eauto. + + rewrite !PArray.get_set_other by auto; auto. +Qed. + +Lemma valid_same_corr : + forall r a1 a2 f1 f2, + ~ is_true (Euf_Checker.valid t_func a2 f2 r) -> + is_true (check_array Atom.eqb a1 a2) -> + is_true (check_array form_eqb f1 f2) -> + is_true (check_local_form (PArray.length a1) f1) -> + is_true (Atom.wf a1) -> + Forall (fun x => Int63.to_Z (State.Lit.blit x) < Int63.to_Z (PArray.length f1)) + (Misc.to_list r) -> + ~ is_true (Euf_Checker.valid t_func a1 f1 r). +Proof. + unfold is_true. intros * VAL CHKA CHKF LOCAL WFA LT. + unfold not in *. intros * VAL2; apply VAL. + unfold Euf_Checker.valid in *. rewrite Misc.afold_left_and in *. + rewrite forallb_forall in *; intros x IN. + pose proof IN as VAL3. + apply VAL2 in VAL3. + rewrite Forall_forall in LT. + pose proof IN as VAL4. + apply LT in VAL4. + unfold interp_state_var. + unfold State.Lit.interp. + unfold State.Var.interp. + pose proof CHKF as INTERP. + eapply t_interp_check_form in INTERP; eauto. + rewrite <- INTERP. + now rewrite get_t_interp_form with (a1 := a1). +Qed. + +End PROOF. + +Lemma restrict_a : + forall p a a', + (forall y, (y <= max_hash_pred p)%positive -> a ! y = a' ! y) -> + eval_hash_pred p a = eval_hash_pred p a'. +Proof. + induction p; crush. + - now rewrite H by lia. + - erewrite IHp1. erewrite IHp2. eauto. + intros. eapply H. lia. + intros. eapply H. lia. + - erewrite IHp1. erewrite IHp2. eauto. + intros. eapply H. lia. + intros. eapply H. lia. + - erewrite IHp1. erewrite IHp2. eauto. + intros. eapply H. lia. + intros. eapply H. lia. + - erewrite IHp. eauto. + intros. eapply H. lia. +Qed. + +Definition nuke_asgn p (a: PTree.t bool) := + PTree.fold (fun b i a => + if (i <=? max_hash_pred p)%positive then PTree.set i a b else b + ) a (PTree.empty bool). + +Lemma eval_hash_pred_ext: + forall m m': PTree.t bool, + (forall x : PTree.elt, m ! x = m' ! x) -> + forall p, + eval_hash_pred p m = eval_hash_pred p m'. +Proof. + intros. eapply restrict_a; eauto. +Qed. + +Lemma nuke_asgn_spec : + forall p a, + exists a', + eval_hash_pred p a = eval_hash_pred p a' + /\ (forall y, (max_hash_pred p < y)%positive -> a' ! y = None). +Proof. + intros. exists (nuke_asgn p a). + revert p. revert a. + split. + - eapply restrict_a; eauto. + intros. unfold nuke_asgn. + eapply PTree_Properties.fold_rec; eauto. + + intros. rewrite <- H0; auto. + + intros. + case_eq ((k <=? max_hash_pred p)%positive); intros. + * destruct (peq k y). + -- subst. rewrite ! PTree.gss; auto. + -- rewrite ! PTree.gso; auto. + * rewrite POrderedType.Positive_as_DT.leb_gt in H3. + assert (k <> y) by lia. + rewrite PTree.gso; auto. + - unfold nuke_asgn. + eapply PTree_Properties.fold_rec; eauto; intros. + + unfold PTree.empty. rewrite PTree.gleaf; auto. + + destruct_match. + * rewrite POrderedType.Positive_as_DT.leb_le in Heqb. + assert (k <> y) by lia. + rewrite PTree.gso; auto. + * eapply H1; eauto. +Qed. + +Lemma nuke_1 : + forall p a x, + ~ pred_In x p -> + eval_hash_pred p a = eval_hash_pred p (PTree.remove x a). +Proof. + induction p; crush. + - assert (a <> x) by (unfold not; intros; subst; apply H; constructor). + rewrite PTree.gro; auto. + - assert (~ pred_In x p1) by (unfold not; intros; apply H; apply pred_In_Pand1; auto). + assert (~ pred_In x p2) by (unfold not; intros; apply H; apply pred_In_Pand2; auto). + erewrite IHp1; [|eassumption]; erewrite IHp2; eauto. + - assert (~ pred_In x p1) by (unfold not; intros; apply H; apply pred_In_Por1; auto). + assert (~ pred_In x p2) by (unfold not; intros; apply H; apply pred_In_Por2; auto). + erewrite IHp1; [|eassumption]; erewrite IHp2; eauto. + - assert (~ pred_In x p1) by (unfold not; intros; apply H; apply pred_In_Pimp1; auto). + assert (~ pred_In x p2) by (unfold not; intros; apply H; apply pred_In_Pimp2; auto). + erewrite IHp1; [|eassumption]; erewrite IHp2; eauto. + - assert (~ pred_In x p) by (unfold not; intros; apply H; constructor; auto). + erewrite IHp; eauto. +Qed. + +Lemma check_atom_to_smt : + forall a p d forms atoms, + to_smt a p = Some (d, forms, atoms) -> + is_true (check_atom atoms). +Proof. + intros. unfold to_smt in *. repeat (destruct_match; try discriminate; []). + crush. +Qed. + +Lemma check_wt_extra_spec : + forall a, + is_true (check_atom a) -> + is_true (check_wt_extra a) -> + (forall i x l, PArray.get a i <> Anop x l) + /\ (forall i n l, PArray.get a i = Aapp n l -> l = nil). +Proof. + unfold check_wt_extra, is_true; intros. + split; intros. + - unfold not; intros. exploit PArray.get_not_default_lt. rewrite H1. + unfold check_atom in H. destruct (PArray.default a); try discriminate. + intros. eapply Misc.aforallbi_spec in H0; eauto. + rewrite H1 in H0. discriminate. + - unfold not; intros. exploit PArray.get_not_default_lt. rewrite H1. + unfold check_atom in H. destruct (PArray.default a); try discriminate. + intros. eapply Misc.aforallbi_spec in H0; eauto. + rewrite H1 in H0. destruct_match; crush. +Qed. + +Lemma pred_in_correct : + forall p c, + pred_in c p = true -> + pred_In c p. +Proof. + induction p; crush; subst. + - constructor. + - apply orb_prop in H. inv H; auto using pred_In_Pand1, pred_In_Pand2. + - apply orb_prop in H. inv H; auto using pred_In_Por1, pred_In_Por2. + - apply orb_prop in H. inv H; auto using pred_In_Pimp1, pred_In_Pimp2. + - constructor; auto. +Qed. + +Lemma pred_in_complete : + forall p c, + pred_In c p -> + pred_in c p = true. +Proof. + induction p; crush; subst; try solve [inv H]. + - inv H. apply Pos.eqb_refl. + - apply orb_true_intro. inv H; intuition. + - apply orb_true_intro. inv H; intuition. + - apply orb_true_intro. inv H; intuition. + - inv H. auto. +Qed. + +Lemma max_key_upd_asgn : + forall l a x, + Forall (fun y => Z.pos y < x) (map fst l) -> + (Z.pos (max_key a) < x) -> + (Z.pos (max_key (upd_asgn l a)) < x). +Proof. + induction l; crush. + inv H. destruct_match. eapply IHl; auto. + - assert (exists y, (PTree.set (fst a) b a0) ! (max_key (PTree.set (fst a) b a0)) = Some y). + { apply max_key_in. + destruct (PTree.bempty (PTree.set (fst a) b a0)) eqn:?; auto. + exfalso. apply PTree.bempty_correct with (x := (fst a)) in Heqb0. rewrite PTree.gss in Heqb0. discriminate. + } + simplify. destruct (peq (max_key (PTree.set (fst a) b a0)) (fst a)). + + lia. + + rewrite PTree.gso in H1; auto. + assert (forall x y, ~ y <= x -> x < y) by lia. + apply H; unfold not; intros. + assert (Z.pos (max_key a0) < Z.pos (max_key (PTree.set (fst a) b a0))) by lia. + pose proof (max_not_present _ (max_key (PTree.set (fst a) b a0)) a0). + rewrite H6 in H1 by lia. discriminate. + - eapply IHl; eauto. +Qed. + +Theorem valid_check_smt : + forall p a, + check_smt p = true -> eval_hash_pred p a = Some true. +Proof. + intros * CHK. unfold check_smt in CHK. repeat (destruct_match; try discriminate; []). + apply andb_prop in CHK; inversion CHK as [CHK1 H0]; clear CHK. + apply andb_prop in CHK1; inversion CHK1 as [CHK6 CHKROOT]; clear CHK1. + apply andb_prop in CHK6; inversion CHK6 as [CHK2 CHK_EXTRA]; clear CHK6. + apply andb_prop in CHK2; inversion CHK2 as [CHK3 H2]; clear CHK2. + apply andb_prop in CHK3; inversion CHK3 as [CHK4 H3]; clear CHK3. + apply andb_prop in CHK4; inversion CHK4 as [CHK5 H4]; clear CHK4. + apply andb_prop in CHK5; inversion CHK5 as [H H5]; clear CHK5. + apply andb_prop in CHKROOT; inversion CHKROOT as [CHKROOT2 P1LT]; clear CHKROOT. + apply andb_prop in CHKROOT2; inversion CHKROOT2 as [CHKROOT3 NOTPREDIN]; clear CHKROOT2. + simplify. destruct p0. + exploit nuke_asgn_spec. intros. inv H1. inv H6. rewrite H1. + assert (NOTPRED: ~ pred_In 1%positive p). + { unfold not; intros HH. apply pred_in_complete in HH. rewrite HH in NOTPREDIN. discriminate. } + erewrite nuke_1; eauto. + assert (ALLEQ: forall y, (max_hash_pred p < y)%positive -> (PTree.remove 1 x) ! y = None). + { intros. destruct (peq 1 y); subst. rewrite PTree.grs. auto. rewrite PTree.gro; auto. } + erewrite unnest_predicate_correct; try eassumption; try lia; [|intros; eapply ALLEQ; lia]. + pose proof (exists_asgn (upd_asgn (rev l) (PTree.remove 1 x))) as X. + assert (ASSUMP1: Z.pos (max_key (upd_asgn (rev l) (PTree.remove 1 x))) - 2 < wB - 1). + { assert (forall l a x, + Forall (fun y => Z.pos y < x) (map fst l) -> + (Z.pos (max_key a) < x) -> + (Z.pos (max_key (upd_asgn l a)) < x)). + { apply max_key_upd_asgn; auto. } + assert (Z.pos (max_key (upd_asgn (rev l) (PTree.remove 1 x))) < wB). + { assert (GTWB: Z.pos p1 < wB) by assumption. eapply H6. apply Forall_forall; intros. + apply map_in_some in H8. inv H8. + inv H9. destruct x1 in *. apply in_rev in H8. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + cbn [fst snd] in *. lia. + assert ((max_key (PTree.remove 1 x) <= max_hash_pred p)%positive). + { assert (~ (max_hash_pred p < max_key (PTree.remove 1 x))%positive -> (max_key (PTree.remove 1 x) <= max_hash_pred p)%positive) by lia. + apply H8. unfold not; intros. pose proof H9. apply ALLEQ in H9. + destruct (PTree.bempty (PTree.remove 1 x)) eqn:?. + - rewrite max_key_empty in * by auto. lia. + - exploit max_key_in; try eassumption. intros. inv H11. rewrite H12 in H9. discriminate. + } + assert (max_hash_pred p < p1)%positive. + { eapply unnest_predicate_lt_fresh in Heqp0. lia. } + lia. + } + lia. + } + assert (ASSUMP2: (forall (k : positive) (x0 : bool), + (upd_asgn (rev l) (PTree.remove 1 x)) ! k = Some x0 -> 0 <= Z.pos k - 2 < wB)). + { intros. pose proof H6 as XXX. apply max_key_correct in H6. + destruct (in_dec peq k (map fst l)). + { apply map_in_some in i. inv i. inv H8. destruct x1; cbn [fst] in *. + apply unnest_predicate_lt_in_list in Heqp0. eapply Forall_forall in Heqp0; eauto. + cbn [fst] in *. lia. + } + { assert (~ In k (map fst (rev l))). + { unfold not; intros; apply n. rewrite map_rev in H8. apply in_rev in H8. auto. } + rewrite upd_asgn_not_in in XXX by auto. + assert (k = 1%positive \/ (1 < k)%positive) by lia. inv H9. rewrite PTree.grs in XXX. discriminate. + lia. + } + } + assert (X2 := X). clear X. + pose proof (X2 ASSUMP1 ASSUMP2) as X. + inversion X as [res ASGN]; clear X. + assert (ASSUMP3: forall (i : Int63.int) (x0 : nop) (l0 : list Int63.int), + PArray.get a5 i <> Anop x0 l0) by (eapply check_wt_extra_spec; eauto). + assert (ASSUMP4: forall (i n : Int63.int) (l0 : list Int63.int), + PArray.get a5 i = Aapp n l0 -> l0 = nil) by (eapply check_wt_extra_spec; eauto). + eapply wt_correct in H2; try eassumption. + pose proof (Checker_Ext.checker_ext_correct _ _ _ _ _ H0) as X. + assert (XX: is_true (Atom.wf a0)) by + (apply check_atom_to_smt in Heqo; unfold check_atom in *; + destruct (PArray.default a0); try discriminate; + destruct c0; crush). + eapply valid_same_corr in X; try eassumption. + eapply valid_to_smt; try eassumption. + simpl. eapply fold_left_Pand_rev. + rewrite to_equiv_rev. eapply eval_hash_pred_equiv_asgn; auto. + unfold equiv_list_wf, Sorted; split. + replace pl_lt with (Basics.flip (Basics.flip (@pl_lt predicate))) by auto. + eapply sorted_rev4. eapply unnest_predicate_sorted; eauto. + apply Forall_rev; eapply unnest_predicate_lt_in_list2; eauto; lia. + apply Forall_rev; apply Forall_forall; intros. eapply ALLEQ. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp0) as Y. eapply Forall_forall in Y; eauto. lia. + apply Forall_rev; apply Forall_forall; intros. simpl. + pose proof (unnest_predicate_lt_in_list _ _ _ _ _ Heqp0) as Y. eapply Forall_forall in Y; eauto. lia. + apply Forall_forall; intros. unfold check_root_index in CHKROOT3. + eapply forallb_forall in CHKROOT3; eauto. + lia. +Qed. + +Lemma mutate_eval_correct : + forall p a, + eval_hash_pred (mutate1_p p) a = eval_hash_pred p a. +Proof. + unfold mutate1_p; intros. + repeat (destruct_match; crush). +Qed. + +Theorem valid_check_smt_total : + forall p a, + check_smt_total p = true -> eval_hash_pred p a = Some true. +Proof. + unfold check_smt_total; intros. + destruct_match; auto using valid_check_smt. + rewrite <- mutate_eval_correct. + destruct_match; auto using valid_check_smt. + discriminate. +Qed. -- cgit