Require Import mathcomp.ssreflect.ssreflect.
Require Import CTL_def dags.

Set Implicit Arguments.
Import Prenex Implicits.

Implicit Types (C D E F : clause) (S T : {fset clause}) (G : rGraph clause).

Relaxed Demos



Definition AXn C := [fset s in C | isDia s].
Definition Req C := [fset R C] `|` [fset body s |` R C | s <- AXn C].

Lemma in_Req u C : fAX u^- \in C -> u^-|` R C \in Req C.

Lemma inhab_Req C : exists D, D \in Req C.

Lemma Req_RC C D : D \in Req C -> R C `<=` D.

Relaxed Fulfillment

We define the fulfillment relations using fixpoint iteration. This turns out to be technically convenient since we never prove anything by induction on the fulfillment relations

Section Fulfillment.
  Variable S T: {fset clause}.

  Definition fulfillAU_step s t (X : {fset clause}) :=
    [fset Y in T | [all C in Req Y,
       suppS S (t^+ |` C) || [some L in T, suppC L (s^+ |` C) && (L \in X)]]].
  Definition fulfillsAU s t L := L \in fset.lfp T (fulfillAU_step s t).

  Lemma bounded_fulfillAU_step s t : bounded T (fulfillAU_step s t).

  Lemma monotone_fulfillAU_step s t : monotone (fulfillAU_step s t).

  Definition fulfillsAUE s t := fset.lfpE (monotone_fulfillAU_step s t) (bounded_fulfillAU_step s t).

  Definition fulfillAR_step s t (X : {fset clause}) :=
    [fset Y in T |
     [all D in Req Y, suppS S D] &&
      (suppS S (t^- |` R Y) || [some E in T, suppC E (s^- |` R Y) && (E \in X)])].
  Definition fulfillsAR s t L := L \in fset.lfp T (fulfillAR_step s t).

  Lemma bounded_fulfillAR_step s t : bounded T (fulfillAR_step s t).

  Lemma monotone_fulfillAR_step s t : monotone (fulfillAR_step s t).

  Definition fulfillsARE s t := fset.lfpE (monotone_fulfillAR_step s t) (bounded_fulfillAR_step s t).

End Fulfillment.

Definition AXn_complete (S : {fset clause}) :=
  forall L C, L \in S -> C \in Req L -> exists L', (L' \in S) && suppC L' C.

Record demo (S T : {fset clause}) :=
  Demo { demoT : forall C, C \in T -> lcons C;
         demoSsubT : {subset S <= T};
         demo1 : AXn_complete S;
         demo2 s t L : L \in S -> fAX (fAU s t)^+ \in L -> fulfillsAU S T s t L;
         demo3 s t L : L \in S -> fAX (fAR s t)^- \in L -> fulfillsAR S T s t L}.

Fragment Demos


Section FragmentDemo.
  Variables (S T : {fset clause}).

  Definition Fs : {fset sform} := \bigcup_(C in S) C.

  Definition req_cond G :=
    forall x : G, ~~ leaf x -> forall D, D \in Req (label x) -> exists2 y, edge x y & suppC (label y) D.

  Definition supp_cond G :=
    forall (x y : G), edge x y -> suppC (label y) (R (label x)).

  Definition ev_cond u G :=
    let C := label (root G) in
       match u with
         | fAX (fAU s t)^+ =>
           (u \in C) && C |> s^+ -> forall x:G, if leaf x then label x |> t^+ else label x |> s^+
         | fAX (fAR s t)^- =>
           (u \in C) && C |> s^- -> (forall x:G, ~~ leaf x -> label x |> s^-)
                                 /\ (exists x:G, label x |> t^-)
         | _ => True
       end.

  Definition glabels_from (G : graph clause) :=
    forall x : G, if leaf x then label x \in S else label x \in T.

  Record fragment G u L := {
    DG1 : label (root G) = L;
    DG2 : terminates (erel G);
    DG3 : req_cond G;
    DG4 : supp_cond G;
    DG5 : ev_cond u G;
    DG6 : ~~ leaf (root G);
    DG7 : glabels_from G
  }.

  Record fragment_demo :=
    FragmentDemo {
        fd_trees L u :> (u \in Fs) -> (L \in S) -> { G : rGraph clause | fragment G u L };
        fd_sub : {subset S <= T};
        fd_lcons C : C \in T -> lcons C
      }.

  Definition fragment_bound (DD : fragment_demo) n :=
    forall L u (inFs : u \in Fs) (inD : L \in S), #|sval (DD _ _ inFs inD)| <= n.
End FragmentDemo.

Model Construction

The model construction for CTL consists of two parts: the construction of fragments form inductive fulfillment and the assemly of fragments into a model.

Relaxed Fullfillment to Fragments

We now construct the fragments required for the model constructioon. There are three cases:
  • Fragments for fAX (fAU s t)^+ \in C
  • Fragments for fAX (fAR s t)^- \in C
  • Fragments where the eventuallity condition is trivial
This case split is justfied by the lemma below.

Inductive ev_case (C : clause) (s : sform) : Type :=
  | ev_AR u v : s = fAX (fAR u v)^- -> C |> u^- -> C |> fAX (fAR u v)^- -> ev_case C s
  | ev_AU u v : s = fAX (fAU u v)^+ -> C |> u^+ -> C |> fAX (fAU u v)^+ -> ev_case C s
  | ev_x : (forall G, label (root G) = C -> ev_cond s G) -> ev_case C s.

Lemma evP C s : ev_case C s.

Section DemoToFragment.
  Variables (S T : {fset clause}) (demoST : demo S T).

  Section Fragments.
  Variables (s t : form).

Fragments for AU
The construction of fragments for AU is completely declarative (no recursion involved). Even the verification of the fragment conditions does not use induction

  Definition AU_level C (rC : fulfillsAU S T s t C) : nat := levl (monotone_fulfillAU_step S T s t) rC.

  Definition AUl := [fset C in T | fulfillsAU S T s t C].
  Definition AUr := [fset C in S | C |> t^+].

  Lemma AUlE C : C \in AUl -> fulfillsAU S T s t C.

  Definition AU_T0 := (seq_sub AUl + seq_sub AUr)%type.
  Definition AU_label (x : AU_T0) := match x with inl x => val x | inr x => val x end.

  Definition AU_rel (x y : AU_T0) :=
    match x,y with
     | inl (SeqSub C rC),inl (SeqSub D rD) =>
       [&& D |> s^+, suppC D (R C) & (AU_level (AUlE rD) < AU_level (AUlE rC))]
     | inl (SeqSub C rC),inr (SeqSub D rD) => suppC D (R C)
     | inr _,_ => false
    end.

  Lemma AU_rel_succ (x : seq_sub AUl) D :
    D \in Req (ssval x) -> exists2 y, AU_rel (inl x) y & suppC (AU_label y) D.

  Lemma AU_rel_term : terminates AU_rel.

  Lemma AU_rel_supp x y : AU_rel x (inl y) -> AU_label (inl y) |> s^+.

  Lemma frag_AU C : C |> s^+ -> fAX (fAU s t)^+ \in C -> fAX (fAU s t)^+ \in Fs S -> C \in S ->
    {G : rGraph clause | fragment S T G (fAX (fAU s t)^+) C & #|G| <= 2 * size T}.

Fragments for AR
The construction of fragments for AR is also completely declarative. Unlike for AU, the verification of the fragment conditions requires an induction the level to locate the clause fulfilling t^-

  Definition AR_level C (rC : fulfillsAR S T s t C) : nat := levl (monotone_fulfillAR_step S T s t) rC.

  Definition ARl := [fset C in T | fulfillsAR S T s t C].
  Definition ARr := S.

  Lemma ARlE C : C \in ARl -> fulfillsAR S T s t C.

  Definition AR_T0 := (seq_sub ARl + seq_sub ARr)%type.

  Definition AR_rel (x y : AR_T0) :=
    match x,y with
      | inl (SeqSub C rC),inl (SeqSub D rD) =>
        [&& D |> s^-, suppC D (R C) & (AR_level (ARlE rD) < AR_level (ARlE rC))]
      | inl (SeqSub C rC),inr (SeqSub D rD) => suppC D (R C)
      | inr _,_ => false
    end.

  Definition AR_label (x : AR_T0) := match x with inl x => val x | inr x => val x end.

  Lemma AR_rel_succ (x : seq_sub ARl) D :
    D \in Req (ssval x) -> exists2 y, AR_rel (inl x) y & suppC (AR_label y) D.

  Lemma AR_rel_term : terminates AR_rel.

  Lemma AR_ev_cond (x : seq_sub ARl) :
    exists2 y, connect AR_rel (inl x) y & AR_label y |> t^-.

  Lemma AR_rel_supp x y : AR_rel x (inl y) -> AR_label (inl y) |> s^-.

  Lemma frag_AR C : C |> s^- -> fAX (fAR s t)^- \in C -> fAX (fAR s t)^- \in Fs S -> C \in S ->
    {G : rGraph clause | fragment S T G (fAX (fAR s t)^-) C & #|G| <= 2 * size T}.

  End Fragments.

  Lemma simple_frag s C : (forall G, label (root G) = C -> ev_cond s G) ->
    C \in S -> { G : rGraph clause | fragment S T G s C & #|G| <= 2 * size T}.

We package the fragments into a fragment demp. Save for the size-bound for the constructed model, the construction of models from fragments is independent of the size of the fragments. Hence, we keep the fragment_demo separate from the ragment_bound.

  Definition bounded_demo (D Dl : {fset clause}) (n : nat) :=
    forall L u (inFs : u \in Fs D) (inD : L \in D), {G : rGraph clause | fragment D Dl G u L & #|G| <= n}.

  Lemma DD_bounded_demo : bounded_demo S T (2 * size T).

  Lemma DD_demo_aux : { FRAG : fragment_demo S T | fragment_bound FRAG (2 * size T) }.

  Definition DD_demo : fragment_demo S T := sval DD_demo_aux.

  Lemma DD_fragment_bound : fragment_bound DD_demo (2 * size T).

End DemoToFragment.

Fragments to Models

interp' is convertible to the sltype.interp with the slpType instance for formulas given in hilbert.v. The local definition avoids a spurious dependency of the model construction on the Hilbert system

Definition interp' s := match s with u^+ => u | u^- => fImp u fF end.

Module ModelConstruction.
Section ModelConstruction.
  Variables (Dl D : {fset clause}) (FD : fragment_demo D Dl).

We turn both Fs D and D into finite types and construct a finite matrix of fragments indexed by the finite type DF * DC

  Definition DF := seq_sub (Fs D).
  Definition DC := seq_sub D.

  Definition Frag (p: DF * DC) := (sval (FD.(fd_trees) (ssvalP p.1) (ssvalP p.2))).

  Lemma FragP (p: DF * DC) : fragment D Dl (Frag p) (val p.1) (val p.2).

  Definition MType := { p : DF * DC & Frag p }.
  Definition MLabel p (x:MType) := fV p^+ \in label (tagged x).

For the definition of the transistion relation, we use the litf_edge construction to lift all the internal edges of the various dags to transitions in the model. We define a cyclic next fuction we define on finite types

  Definition MRel (x y : MType) :=
     lift_edge x y ||
      (let: (existT p a, (existT q a')) := (x,y) in
       [&& leaf a, q.1 == next p.1, edge (root (Frag q)) a' & label a == val q.2 ]).

Cast from nodes of fragments fo states of the model
  Definition Mstate (p: DF * DC) (x : Frag p) : MType := existT p x.


To construct a finite model, we need to show that MRel is serial

  Lemma label_DF (p : DF * DC) (x : Frag p) : leaf x -> label x \in D.

  Lemma root_internal (p : DF * DC) : ~~ leaf (root (Frag p)).

  Lemma serial_MRel (x : MType) : exists y, MRel x y.

  Definition MD := FModel MLabel serial_MRel.

We prove a number of lemmas to extract fragment properties and lift porperies from the dags to the model MD

  Lemma label_root (p : DF * DC) : label (root (Frag p)) = val (p.2).

  Lemma dmat_ev_cond (p : DF * DC) : ev_cond (val p.1) (Frag p).

  Lemma sn_edge (p : DF * DC) (x : Frag p) : sn (edge (g := Frag p)) x.

  Lemma label_Dl (p : DF * DC) (x : Frag p) : ~~ leaf x -> label x \in Dl.

  Lemma label_lcon (x : MType) : lcons (label (tagged x)).

  Lemma rlabel_Fs p : {subset label (root (Frag p)) <= Fs D}.

  Lemma dag_R (p : DF * DC) (x y : Frag p) : edge x y -> suppC (label y) (R (label x)).

  Lemma MRel_R (x y : MType) : MRel x y -> suppC (label (tagged y)) (R (label (tagged x))).

  Lemma dag_D (p : DF * DC) (x : Frag p) s :
    ~~ leaf x -> fAX s^- \in label x -> exists2 y, edge x y & label y |> s^-.

  Lemma MRel_D (x : MType) s :
    fAX s^- \in label (tagged x) -> exists2 y, MRel x y & label (tagged y) |> s^-.

  Lemma supp_edge (p : DF * DC) (x y : Frag p) s : label x |> fAX s^+ -> edge x y -> label y |> s^+.


We define special predicates over MD that express the eventualities in terms of support instead of eval. In the proof of the support lemma, this will be what is left to show after the induction hypothesis has been applied.

  Definition AU_M s t (x : MD) :=
    cAU MRel (fun z => label (tagged z) |> s^+) (fun z => label (tagged z) |> t^+) x.

  Definition EU_M s t (x : MD) :=
    cEU MRel (fun z => label (tagged z) |> s^-) (fun z => label (tagged z) |> t^-) x.

We make use uf the fact, that CTL cannot distinguish between equally labeled states that have the same successorts. In particular this means that an eventuallity is satisfied at the leaf of one dag, if it is satiesfied at the equally labeled root of the next row.

  Lemma leaf_root_aux (p : DF * DC) (x : Frag p) (lf : leaf x) :
      forall y, MRel (Mstate (root (Frag (next p.1, SeqSub (label x) (label_DF lf))))) y <->
                MRel (Mstate x) y.

  Lemma leaf_root s t (p : DF * DC) (x : Frag p) (lf : leaf x) :
    AU_M s t (Mstate (root (Frag (next p.1, SeqSub (label x) (label_DF lf))))) ->
    AU_M s t (Mstate x).

  Lemma leaf_root' s t (p : DF * DC) (x : Frag p) (lf : leaf x) :
    EU_M s t (Mstate (root (Frag (next p.1, SeqSub _ (label_DF lf))))) ->
    EU_M s t (Mstate x).

Every eventuality is fulfilled at the root of its correspinding dag.

  Lemma AU_this (s t : form) (L : DC) (inF : fAX (fAU s t)^+ \in Fs D) :
    let p := (SeqSub (fAX (fAU s t)^+) inF, L) in
    label (root (Frag p)) |> fAU s t^+ -> AU_M s t (Mstate (root (Frag p))).

  Lemma EU_this (s t : form) (L : DC) (inF : fAX (fAR s t)^- \in Fs D) :
    let p := (SeqSub (fAX (fAR s t)^-) inF, L) in
    label (root (Frag p)) |> fAR s t^- -> EU_M s t (Mstate (root (Frag p))).

We can always defer fullfilling an eventuality to the next row.

  Lemma EU_lift (s t : form) (p : DF * DC) (x : Frag p) :
    (forall L', label (root (Frag (next p.1, L'))) |> fAR s t^-
             -> EU_M s t (Mstate (root (Frag (next p.1, L'))))) ->
    label x |> fAR s t^- -> EU_M s t (Mstate x).

  Lemma AU_lift (s t : form) (p : DF * DC) (x : Frag p) :
    (forall L', label (root (Frag (next p.1, L'))) |> fAU s t^+
             -> AU_M s t (Mstate (root (Frag (next p.1, L'))))) ->
    label x |> fAU s t^+ -> AU_M s t (Mstate x).

Finally, we can prove that every state in the model satisfies all signed formulas it supports. The interesting cases are those for the eventualities. In either case we first show that it suffices to consider eventualities at root of a fragment by defering to the next row in the matix. Only after this we start the induction on the distance between the current row and the row for the eventuality under consideration

  Lemma supp_eval s (x : MD) : label (tagged x) |> s -> eval (interp' s) x.

  Lemma model_existence u L :
    u \in Fs D -> L \in D -> exists (M : fmodel) (w : M), forall s, L |> s -> eval (interp' s) w.

Stronger form of model existence if there is a known bound on the fragment size
  Lemma MD_size n : fragment_bound FD n -> #|MD| <= n * (size (Fs D) * size D).

  Lemma small_model_existence u L n : fragment_bound FD n -> u \in Fs D -> L \in D ->
     exists2 M : fmodel, #|M| <= n * (size (Fs D) * size D) &
                 exists (w : M), forall s, L |> s -> eval (interp' s) w.

End ModelConstruction.
End ModelConstruction.
Import ModelConstruction.

Variant that emphasizes that all clauses of a demo are satsified by the same model

Theorem small_models S T u (inF : u \in Fs S) :
  demo S T -> exists2 M : fmodel,
   #|M| <= 2 * size (Fs S) * (size T)^2 &
   forall C, C \in S -> exists (w : M), forall s, C |> s -> eval (interp' s) w.

Unit Model

Definition relT {X:Type} := [rel x y : X | true].

Lemma serial_relT (X:Type) : forall x: X, exists y, relT x y.

Definition unit_model := FModel (fun _ => xpred0) (@serial_relT unit).