Library OracleLanguages.Imp.Oracles.SwapBranches


Require Import Common.
Require Import Imp.Lang.
Require Import ProgrammingLanguage.
Require Import OracleLanguage.
Require Import OracleHelpers.

Module Make
       (Var_as_DT : UsualDecidableTypeOrig)
       (Import Lang : ImpProgrammingLanguage Var_as_DT)
       (Import OracleHelpers : ImpOracleHelpers Var_as_DT Lang).
  Inductive mod :=
  | Negate | UnNegate.

  Inductive _cont_mod := .

  Definition _genv := unit.

  Inductive _state : Type :=
  | State : state state list (cont_mod mod _cont_mod) _state.

  Definition _step (_ : _genv) (S : _state) : option _state :=
    let '(State S₁ S₂ km) := S in
    step_helper S₁ S₂ km >>= λ sr,
    match sr with
    | StuckStep
      None
    | GenericStep S₁' S₂' km'
      Some (State S₁' S₂' km')
    | SpecialCmdStep _ km'
      match step υ S₁, step υ S₂ with
      | Some S₁', Some S₂'
        Some (State S₁' S₂' (CmdMod Id::km'))
      | _, _
        None
      end
    | SpecialContStep _ km'
      None
    end.

  Definition _left_state (S : _state) : ProgrammingLanguage.state imp_language :=
    match S with
    | State S₁ _ _
      S₁
    end.

  Definition _right_state (S : _state) : ProgrammingLanguage.state imp_language :=
    match S with
    | State _ S₂ _
      S₂
    end.

  Definition syntactic_negate c :=
    match c with
      | ITE b c₁ c₂
        Some (ITE (NOT b) c₂ c₁)
      | _None
    end.

  Definition syntactic_un_negate c :=
    match c with
      | ITE (NOT b) c₁ c₂
        Some (ITE b c₂ c₁)
      | _None
    end.

  Fixpoint apply_cmd_mod c m :=
    match m with
    | IdSome c
    | RecMod m₁ m₂
      match c with
      | Seq c₁ c₂
        apply_cmd_mod c₁ m₁ >>= λ c₁',
        apply_cmd_mod c₂ m₂ >>= λ c₂',
        Some (Seq c₁' c₂')
      | ITE b c₁ c₂
        apply_cmd_mod c₁ m₁ >>= λ c₁',
        apply_cmd_mod c₂ m₂ >>= λ c₂',
        Some (ITE b c₁' c₂')
      | While b c
        apply_cmd_mod c m₁ >>= λ c',
        Some (While b c')
      | Skip | Assign _ _ | Assert _
        None
      end
    | LeafMod Negatesyntactic_negate c
    | LeafMod UnNegatesyntactic_un_negate c
    end.

  Lemma apply_cmd_mod_id:
     c, apply_cmd_mod c Id = Some c.

  Definition leaf_mod_of_cmds (c₁ c₂ : cmd) : option mod :=
    match apply_cmd_mod c₁ (LeafMod Negate) >>= λ c₁',
          if beq_cmd c₁' c₂ then
            Some Negate
          else
            None
    with
    | Some mSome m
    | None
      apply_cmd_mod c₁ (LeafMod UnNegate) >>= λ c₁',
      if beq_cmd c₁' c₂ then
        Some UnNegate
      else
        None
    end.

  Inductive km_invariant : list (cont_mod mod _cont_mod) cont cont Prop :=
  | empty_cont_mod : km_invariant [] [] []
  | cmd_cont : c c' m km k₁ k₂,
      apply_cmd_mod c m = Some c'
      km_invariant km k₁ k₂
      km_invariant (CmdMod m::km) (c::k₁) (c'::k₂).

  Definition _invariant (_ : _genv) (S : _state) : Prop :=
    match S with
    | State (k₁, s₁) (k₂, s₂) km
      km_invariant km k₁ k₂ M.Equal s₁ s₂
    end.

  Lemma _invariant_1:
     genv os os',
      _invariant genv os
      _step genv os = Some os'
      _invariant genv os'.

  Lemma _prediction_soundness :
     genv os os',
      _invariant genv os
      _step genv os = Some os'
       n₁ n₂,
        opt_state_eq (nsteps υ n₁ (_left_state os))
                       (Some (_left_state os'))
         opt_state_eq (nsteps υ n₂ (_right_state os))
                       (Some (_right_state os'))
         n₁ + n₂ > 0.

  Lemma _prediction_completeness :
     genv os,
      _invariant genv os
      _step genv os = None
      step υ (_left_state os) = None
       step υ (_right_state os) = None.

  Definition swapbranches_oracle : oracle_language imp_language imp_language :=
    {|
      oracle_state := _state;
      oracle_genv := unit;
      oracle_step := _step;
      left_state := _left_state;
      right_state := _right_state;
      left_genv := λ _, υ;
      right_genv := λ _, υ;
      invariant := _invariant;
      invariant_1 := _invariant_1;
      prediction_soundness := _prediction_soundness;
      prediction_completeness := _prediction_completeness
    |}.
End Make.