Fsub.ExtendedBase.Fsub_LetSum_Definitions

Definition of Fsub (System F with subtyping).
Authors: Brian Aydemir and Arthur Chargu\'eraud, with help from Aaron Bohannon, Jeffrey Vaughan, and Dimitrios Vytiniotis.
Table of contents:

Require Export ExtendedBase.LabelMap.
Require Export ExtendedBase.Label.
Require Export FqMeta.Metatheory.
Require Export String.
Require Export Lattice.

Here, we are assuming our qualifier lattice
Parameter B : Set.
Context `{O : Order B}.
Context `{L : Lattice B}.
Context `{@LOSet B O L}.

(* ********************************************************************** *)

Syntax (pre-terms)

We use a locally nameless representation for Fsub, where bound variables are represented as natural numbers (de Bruijn indices) and free variables are represented as atoms. The type atom, defined in the MetatheoryAtom library, represents names: there are infinitely many atoms, equality is decidable on atoms, and it is possible to generate an atom fresh for any given finite set of atoms.
We say that the definitions below define pre-types (typ) and pre-expressions (exp), collectively pre-terms, since the datatypes admit terms, such as (typ_all typ_top (typ_bvar 3)), where indices are unbound. A term is locally closed when it contains no unbound indices.
Similarly, in addition to pre-types and pre-expressions, we define pre-qualifiers for representing qualifier terms, with possibly unbound indices.
Note that indices for bound type variables are distinct from indices for bound expression variables. We make this explicit in the definitions below of the opening operations.
A qua, a qualifier term, consists of either a base qualifier lattice element, variables, meets, and joins.
Inductive qua : Set :=
  | qua_base : B qua
  | qua_bvar : qua
  | qua_fvar : atom qua
  | qua_meet : qua qua qua
  | qua_join : qua qua qua
.

Inductive typ : Set :=
  | typ_top : typ
  | typ_bvar : typ
  | typ_fvar : atom typ
  | typ_arrow : qtyp qtyp typ
  | typ_all : typ qtyp typ
  | typ_qall : qua qtyp typ
  | typ_sum : qtyp qtyp typ
  | typ_pair : qtyp qtyp typ
  | typ_ref : qtyp typ
with qtyp : Set :=
  | qtyp_qtyp : qua typ qtyp
.

Inductive exp : Set :=
  | exp_bvar : exp
  | exp_fvar : atom exp
  | exp_abs : qua qtyp exp exp
  | exp_app : exp exp exp
  | exp_tabs : qua typ exp exp
  | exp_tapp : exp typ exp
  | exp_qabs : qua qua exp exp
  | exp_qapp : exp qua exp
  | exp_let : exp exp exp
  | exp_inl : qua exp exp
  | exp_inr : qua exp exp
  | exp_case : exp exp exp exp
  | exp_pair : qua exp exp exp
  | exp_first : exp exp
  | exp_second : exp exp
  | exp_ref : qua exp exp
  | exp_ref_label : qua label exp
  | exp_deref : exp exp
  | exp_set_ref : exp exp exp
  | exp_upqual : qua exp exp
  | exp_check : qua exp exp
.

We declare the constructors for indices and variables to be coercions. For example, if Coq sees a where it expects an exp, it will implicitly insert an application of exp_bvar; similar behavior happens for atoms. Thus, we may write (exp_abs typ_top (exp_app 0 x)) instead of (exp_abs typ_top (exp_app (exp_bvar 0) (exp_fvar x))).

Coercion typ_bvar : nat >-> typ.
Coercion typ_fvar : atom >-> typ.
Coercion exp_bvar : nat >-> exp.
Coercion exp_fvar : atom >-> exp.
Coercion qua_bvar : nat >-> qua.
Coercion qua_fvar : atom >-> qua.

(* ********************************************************************** *)

Concrete Qualifiers

A concrete qualifier, at runtime, is an element of the base lattice.
Definition concrete_qua : Set := B.

concretize partially evaluates a qua into a concrete_qua.
Fixpoint concretize (q : qua) : (option concrete_qua) :=
  match q with
  | qua_base b Some b
  | qua_fvar X None
  | qua_bvar n None
  | qua_join
    match (concretize , concretize ) with
    | (Some , Some ) Some ( )
    | _ None
    end
  | qua_meet
    match (concretize , concretize ) with
    | (Some , Some ) Some ( )
    | _ None
    end
  end.

abstractize goes the other way, giving us a qualifier term from a concrete, runtime qualifier.
Definition abstractize (s : concrete_qua) :=
  qua_base s.
#[export] Hint Transparent abstractize : core.

Unlike before, B carries a linear order <=, so we don't need to define our own.

(* ********************************************************************** *)

Opening terms

Opening replaces an index with a term. This operation is required if we wish to work only with locally closed terms when going under binders (e.g., the typing rule for exp_abs). It also corresponds to informal substitution for a bound variable, which occurs in the rule for beta reduction.
We need to define three functions for opening due the syntax of Fsub, and we name them according to the following convention.
  • tt: Denotes an operation involving types appearing in types.
  • te: Denotes an operation involving types appearing in expressions.
  • ee: Denotes an operation involving expressions appearing in expressions.
The notation used below for decidable equality on natural numbers (e.g., K == J) is defined in the CoqEqDec library, which is included by the Metatheory library. The order of arguments to each "open" function is the same. For example, (open_tt_rec K U T) can be read as "substitute U for index K in T."
Note that we assume U is locally closed (and similarly for the other opening functions). This assumption simplifies the implementations of opening by letting us avoid shifting. Since bound variables are indices, there is no need to rename variables to avoid capture. Finally, we assume that these functions are initially called with index zero and that zero is the only unbound index in the term. This eliminates the need to possibly subtract one in the case of indices.

Fixpoint open_qq_rec (K : ) (R : qua) (Q : qua) {struct Q} : qua :=
  match Q with
  | qua_base b qua_base b
  | qua_bvar J if (K == J) then R else (qua_bvar J)
  | qua_fvar X qua_fvar X
  | qua_meet qua_meet (open_qq_rec K R ) (open_qq_rec K R )
  | qua_join qua_join (open_qq_rec K R ) (open_qq_rec K R )
  end.

Fixpoint open_tt_rec (K : ) (U : typ) (T : typ) {struct T} : typ :=
  match T with
  | typ_top typ_top
  | typ_bvar J if K == J then U else (typ_bvar J)
  | typ_fvar X typ_fvar X
  | typ_arrow typ_arrow (open_tqt_rec K U ) (open_tqt_rec (S K) U )
  | typ_all typ_all (open_tt_rec K U ) (open_tqt_rec (S K) U )
  | typ_qall Q T typ_qall Q (open_tqt_rec (S K) U T)
  | typ_sum typ_sum (open_tqt_rec K U ) (open_tqt_rec K U )
  | typ_pair typ_pair (open_tqt_rec K U ) (open_tqt_rec K U )
  | typ_ref typ_ref (open_tqt_rec K U )
  end
with open_tqt_rec (K : ) (U : typ) (T : qtyp) {struct T} : qtyp :=
  match T with
  | qtyp_qtyp Q T qtyp_qtyp Q (open_tt_rec K U T)
  end.

Fixpoint open_qt_rec (K : ) (R : qua) (T : typ) {struct T} : typ :=
  match T with
  | typ_top typ_top
  | typ_bvar J typ_bvar J
  | typ_fvar X typ_fvar X
  | typ_arrow typ_arrow (open_qqt_rec K R ) (open_qqt_rec (S K) R )
  | typ_all typ_all (open_qt_rec K R ) (open_qqt_rec (S K) R )
  | typ_qall Q T typ_qall (open_qq_rec K R Q) (open_qqt_rec (S K) R T)
  | typ_sum typ_sum (open_qqt_rec K R ) (open_qqt_rec K R )
  | typ_pair typ_pair (open_qqt_rec K R ) (open_qqt_rec K R )
  | typ_ref typ_ref (open_qqt_rec K R )
  end
with open_qqt_rec (K : ) (R : qua) (T : qtyp) {struct T} : qtyp :=
  match T with
  | qtyp_qtyp Q T qtyp_qtyp (open_qq_rec K R Q) (open_qt_rec K R T)
  end.

Fixpoint open_te_rec (K : ) (U : typ) (e : exp) {struct e} : exp :=
  match e with
  | exp_bvar i exp_bvar i
  | exp_fvar x exp_fvar x
  | exp_abs P V exp_abs P (open_tqt_rec K U V) (open_te_rec (S K) U )
  | exp_app exp_app (open_te_rec K U ) (open_te_rec K U )
  | exp_tabs P V exp_tabs P (open_tt_rec K U V) (open_te_rec (S K) U )
  | exp_tapp V exp_tapp (open_te_rec K U ) (open_tt_rec K U V)
  | exp_qabs P Q exp_qabs P Q (open_te_rec (S K) U )
  | exp_qapp Q exp_qapp (open_te_rec K U ) Q
  | exp_let exp_let (open_te_rec K U ) (open_te_rec (S K) U )
  | exp_inl P exp_inl P (open_te_rec K U )
  | exp_inr P exp_inr P (open_te_rec K U )
  | exp_case
      exp_case (open_te_rec K U ) (open_te_rec (S K) U ) (open_te_rec (S K) U )
  | exp_pair R exp_pair R (open_te_rec K U ) (open_te_rec K U )
  | exp_first exp_first (open_te_rec K U )
  | exp_second exp_second (open_te_rec K U )
  | exp_ref P exp_ref P (open_te_rec K U )
  | exp_ref_label P l exp_ref_label P l
  | exp_deref exp_deref (open_te_rec K U )
  | exp_set_ref exp_set_ref (open_te_rec K U ) (open_te_rec K U )
  | exp_check P exp_check P (open_te_rec K U )
  | exp_upqual P exp_upqual P (open_te_rec K U )
  end.

Fixpoint open_qe_rec (K : ) (R : qua) (e : exp) {struct e} : exp :=
  match e with
  | exp_bvar i exp_bvar i
  | exp_fvar x exp_fvar x
  | exp_abs P V exp_abs (open_qq_rec K R P) (open_qqt_rec K R V) (open_qe_rec (S K) R )
  | exp_app exp_app (open_qe_rec K R ) (open_qe_rec K R )
  | exp_tabs P V exp_tabs (open_qq_rec K R P) (open_qt_rec K R V) (open_qe_rec (S K) R )
  | exp_tapp V exp_tapp (open_qe_rec K R ) (open_qt_rec K R V)
  | exp_qabs P Q exp_qabs (open_qq_rec K R P) (open_qq_rec K R Q) (open_qe_rec (S K) R )
  | exp_qapp Q exp_qapp (open_qe_rec K R ) (open_qq_rec K R Q)
  | exp_let exp_let (open_qe_rec K R ) (open_qe_rec (S K) R )
  | exp_inl P exp_inl (open_qq_rec K R P) (open_qe_rec K R )
  | exp_inr P exp_inr (open_qq_rec K R P) (open_qe_rec K R )
  | exp_case
      exp_case (open_qe_rec K R ) (open_qe_rec (S K) R ) (open_qe_rec (S K) R )
  | exp_pair P exp_pair (open_qq_rec K R P) (open_qe_rec K R ) (open_qe_rec K R )
  | exp_first exp_first (open_qe_rec K R )
  | exp_second exp_second (open_qe_rec K R )
  | exp_ref P exp_ref (open_qq_rec K R P) (open_qe_rec K R )
  | exp_ref_label P l exp_ref_label (open_qq_rec K R P) l
  | exp_deref exp_deref (open_qe_rec K R )
  | exp_set_ref exp_set_ref (open_qe_rec K R ) (open_qe_rec K R )
  | exp_check P exp_check (open_qq_rec K R P) (open_qe_rec K R )
  | exp_upqual P exp_upqual (open_qq_rec K R P) (open_qe_rec K R )
  end.

Fixpoint open_ee_rec (k : ) (f : exp) (e : exp) {struct e} : exp :=
  match e with
  | exp_bvar i if k == i then f else (exp_bvar i)
  | exp_fvar x exp_fvar x
  | exp_abs P V exp_abs P V (open_ee_rec (S k) f )
  | exp_app exp_app (open_ee_rec k f ) (open_ee_rec k f )
  | exp_tabs P V exp_tabs P V (open_ee_rec (S k) f )
  | exp_tapp V exp_tapp (open_ee_rec k f ) V
  | exp_qabs P Q exp_qabs P Q (open_ee_rec (S k) f )
  | exp_qapp Q exp_qapp (open_ee_rec k f ) Q
  | exp_let exp_let (open_ee_rec k f ) (open_ee_rec (S k) f )
  | exp_inl P exp_inl P (open_ee_rec k f )
  | exp_inr P exp_inr P (open_ee_rec k f )
  | exp_case
      exp_case (open_ee_rec k f )
               (open_ee_rec (S k) f )
               (open_ee_rec (S k) f )
  | exp_pair P exp_pair P (open_ee_rec k f ) (open_ee_rec k f )
  | exp_first exp_first (open_ee_rec k f )
  | exp_second exp_second (open_ee_rec k f )
  | exp_ref P exp_ref P (open_ee_rec k f )
  | exp_ref_label P l exp_ref_label P l
  | exp_deref exp_deref (open_ee_rec k f )
  | exp_set_ref exp_set_ref (open_ee_rec k f ) (open_ee_rec k f )
  | exp_check P exp_check P (open_ee_rec k f )
  | exp_upqual P exp_upqual P (open_ee_rec k f )
  end.

Many common applications of opening replace index zero with an expression or variable. The following definitions provide convenient shorthands for such uses. Note that the order of arguments is switched relative to the definitions above. For example, (open_tt T X) can be read as "substitute the variable X for index 0 in T" and "open T with the variable X." Recall that the coercions above let us write X in place of (typ_fvar X), assuming that X is an atom.

Definition open_tt T U := open_tt_rec 0 U T.
Definition open_te e U := open_te_rec 0 U e.
Definition open_ee e1 e2 := open_ee_rec 0 .
Definition open_qq Q R := open_qq_rec 0 R Q.
Definition open_qe e R := open_qe_rec 0 R e.
Definition open_qt T R := open_qt_rec 0 R T.
Definition open_qqt T R := open_qqt_rec 0 R T.
Definition open_tqt T R := open_tqt_rec 0 R T.

(* ********************************************************************** *)

Local closure

Recall that typ and exp define pre-terms; these datatypes admit terms that contain unbound indices. A term is locally closed, or syntactically well-formed, when no indices appearing in it are unbound. The proposition (type T) holds when a type T is locally closed, and (expr e) holds when an expression e is locally closed.
The inductive definitions below formalize local closure such that the resulting induction principles serve as structural induction principles over (locally closed) types and (locally closed) expressions. In particular, unlike the situation with pre-terms, there are no cases for indices. Thus, these induction principles correspond more closely to informal practice than the ones arising from the definitions of pre-terms.
The interesting cases in the inductive definitions below are those that involve binding constructs, e.g., typ_all. Intuitively, to check if the pre-term (typ_all ) is locally closed, we must check that is locally closed and that is locally closed when opened with a variable. However, there is a choice as to how many variables to quantify over. One possibility is to quantify over only one variable ("existential" quantification), as in
  type_all : forall X T1 T2,
      type T1 ->
      type (open_tt T2 X) ->
      type (typ_all T1 T2) .
Or, we could quantify over as many variables as possible ("universal" quantification), as in
  type_all : forall T1 T2,
      type T1 ->
      (forall X : atom, type (open_tt T2 X)) ->
      type (typ_all T1 T2) .
It is possible to show that the resulting relations are equivalent. The former makes it easy to build derivations, while the latter provides a strong induction principle. McKinna and Pollack used both forms of this relation in their work on formalizing Pure Type Systems.
We take a different approach here and use "cofinite" quantification: we quantify over all but finitely many variables. This approach provides a convenient middle ground: we can build derivations reasonably easily and get reasonably strong induction principles. With some work, one can show that the definitions below are equivalent to ones that use existential, and hence also universal, quantification.
Inductive qual : qua Prop :=
  | qual_base : b,
      qual (qua_base b)
  | qual_fvar : (X : atom),
      qual (qua_fvar X)
  | qual_meet : Q1 Q2,
      qual
      qual
      qual (qua_meet )
  | qual_join : Q1 Q2,
      qual
      qual
      qual (qua_join )
.

Inductive type : typ Prop :=
  | type_top :
      type typ_top
  | type_var : X,
      type (typ_fvar X)
  | type_arrow : T1 T2,
      qtype
      qtype
      type (typ_arrow )
  | type_all : L T1 T2,
      type
      ( X : atom, X `notin` L qtype (open_tqt X))
      type (typ_all )
  | type_qall : L Q T,
      qual Q
      ( X : atom, X `notin` L qtype (open_qqt T X))
      type (typ_qall Q T)
  | type_sum : T1 T2,
      qtype
      qtype
      type (typ_sum )
  | type_pair : T1 T2,
      qtype
      qtype
      type (typ_pair )
with qtype : qtyp Prop :=
  | qtype_qtyp : Q T,
      qual Q
      type T
      qtype (qtyp_qtyp Q T)
.

Inductive expr : exp Prop :=
  | expr_var : x,
      expr (exp_fvar x)
  | expr_abs : L P T e1,
      qual P
      qtype T
      ( x : atom, x `notin` L expr (open_ee x))
      expr (exp_abs P T )
  | expr_app : e1 e2,
      expr
      expr
      expr (exp_app )
  | expr_tabs : L P T e1,
      qual P
      type T
      ( X : atom, X `notin` L expr (open_te X))
      expr (exp_tabs P T )
  | expr_tapp : e1 V,
      expr
      type V
      expr (exp_tapp V)
  | expr_qabs : L P Q e1,
      qual P
      qual Q
      ( X : atom, X `notin` L expr (open_qe X))
      expr (exp_qabs P Q )
  | expr_qapp : e1 Q,
      expr
      qual Q
      expr (exp_qapp Q)
  | expr_let : L e1 e2,
      expr
      ( x : atom, x `notin` L expr (open_ee x))
      expr (exp_let )
  | expr_inl : P e1,
      qual P
      expr
      expr (exp_inl P )
  | expr_inr : P e1,
      qual P
      expr
      expr (exp_inr P )
  | expr_case : L e1 e2 e3,
      expr
      ( x : atom, x `notin` L expr (open_ee x))
      ( x : atom, x `notin` L expr (open_ee x))
      expr (exp_case )
  | expr_pair : P e1 e2,
      qual P
      expr
      expr
      expr (exp_pair P )
  | expr_first : e1,
      expr
      expr (exp_first )
  | expr_second : e1,
      expr
      expr (exp_second )
  | expr_upqual : P e1,
      qual P
      expr
      expr (exp_upqual P )
  | expr_check : P e1,
      qual P
      expr
      expr (exp_check P )
.

We also define what it means to be the body of an abstraction, since this simplifies slightly the definition of reduction and subsequent proofs. It is not strictly necessary to make this definition in order to complete the development.

Definition body_e (e : exp) :=
   L, x : atom, x `notin` L expr (open_ee e x).

(* ********************************************************************** *)

Environments

In our presentation of System F with subtyping, we use a single environment for both typing and subtyping assumptions. We formalize environments by representing them as association lists (lists of pairs of keys and values) whose keys are atoms.
The Metatheory and MetatheoryEnv libraries provide functions, predicates, tactics, notations and lemmas that simplify working with environments. They treat environments as lists of type list (atom * A). The notation (x ~ a) denotes a list consisting of a single binding from x to a.
Since environments map atoms, the type A should encode whether a particular binding is a typing or subtyping assumption. Thus, we instantiate A with the type binding, defined below.

Inductive binding : Set :=
  | bind_sub : typ binding
  | bind_typ : qtyp binding
  | bind_qua : qua binding.

A binding (X ~ bind_sub T) records that a type variable X is a subtype of T, and a binding (x ~ bind_typ U) records that an expression variable x has type U.
We define an abbreviation env for the type of environments, and an abbreviation empty for the empty environment.
Note: Each instance of Notation below defines an abbreviation since the left-hand side consists of a single identifier that is not in quotes. These abbreviations are used for both parsing (the left-hand side is equivalent to the right-hand side in all contexts) and printing (the right-hand side is pretty-printed as the left-hand side). Since nil is normally a polymorphic constructor whose type argument is implicit, we prefix the name with "@" to signal to Coq that we are going to supply arguments to nil explicitly.

Notation env := (list (atom * binding)).
Notation empty := (@nil (atom * binding)).

Examples: We use a convention where environments are never built using a cons operation ((x, a) :: E) where E is non-nil. This makes the shape of environments more uniform and saves us from excessive fiddling with the shapes of environments. For example, Coq's tactics sometimes distinguish between consing on a new binding and prepending a one element list, even though the two operations are convertible with each other.
Consider the following environments written in informal notation.
  1. (empty environment)
  2. x : T
  3. x : T, Y <: S
  4. E, x : T, F
In the third example, we have an environment that binds an expression variable x to T and a type variable Y to S. In Coq, we would write these environments as follows.
  1. empty
  2. x ~ bind_typ T
  3. Y ~ bind_sub S ++ x ~ bind_typ T
  4. F ++ x ~ bind_typ T ++ E
The symbol "++" denotes list concatenation and associates to the right. (That notation is defined in Coq's List library.) Note that in Coq, environments grow on the left, since that is where the head of a list is.

(* ********************************************************************** *)

Well-formedness

A type T is well-formed with respect to an environment E, denoted (wf_typ E T), when T is locally-closed and its free variables are bound in E. We need this relation in order to restrict the subtyping and typing relations, defined below, to contain only well-formed types. (This relation is missing in the original statement of the POPLmark Challenge.)
Note: It is tempting to define the premise of wf_typ_var as (X `in` dom E), since that makes the rule easier to apply (no need to guess an instantiation for U). Unfortunately, this is incorrect. We need to check that X is bound as a type-variable, not an expression-variable; (dom E) does not distinguish between the two kinds of bindings.

Inductive wf_qua : env qua Prop :=
  | wf_qua_base : E b,
      wf_qua E (qua_base b)
  | wf_qua_fvar : R E (X : atom),
      binds X (bind_qua R) E
      wf_qua E (qua_fvar X)
  | wf_qua_meet : E Q1 Q2,
      wf_qua E
      wf_qua E
      wf_qua E (qua_meet )
  | wf_qua_join : E Q1 Q2,
      wf_qua E
      wf_qua E
      wf_qua E (qua_join )
.

Inductive wf_typ : env typ Prop :=
  | wf_typ_top : E,
      wf_typ E typ_top
  | wf_typ_var : U E (X : atom),
      binds X (bind_sub U) E
      wf_typ E (typ_fvar X)
  | wf_typ_arrow : E T1 T2,
      wf_qtyp E
      wf_qtyp E
      wf_typ E (typ_arrow )
  | wf_typ_all : L E T1 T2,
      wf_typ E
      ( X : atom, X `notin` L
        wf_qtyp (X bind_sub E) (open_tqt X))
      wf_typ E (typ_all )
  | wf_typ_qall : L E T1 T2,
      wf_qua E
      ( X : atom, X `notin` L
        wf_qtyp (X bind_qua E) (open_qqt X))
      wf_typ E (typ_qall )
  | wf_typ_sum : E T1 T2,
      wf_qtyp E
      wf_qtyp E
      wf_typ E (typ_sum )
  | wf_typ_pair : E T1 T2,
      wf_qtyp E
      wf_qtyp E
      wf_typ E (typ_pair )
with wf_qtyp : env qtyp Prop :=
  | wf_qtyp_qtyp : E Q T,
      wf_qua E Q
      wf_typ E T
      wf_qtyp E (qtyp_qtyp Q T)
.

An environment E is well-formed, denoted (wf_env E), if each atom is bound at most at once and if each binding is to a well-formed type. This is a stronger relation than the uniq relation defined in the MetatheoryEnv library. We need this relation in order to restrict the subtyping and typing relations, defined below, to contain only well-formed environments. (This relation is missing in the original statement of the POPLmark Challenge.)

Inductive wf_env : env Prop :=
  | wf_env_empty :
      wf_env empty
  | wf_env_sub : (E : env) (X : atom) (T : typ),
      wf_env E
      wf_typ E T
      X `notin` dom E
      wf_env (X bind_sub T E)
  | wf_env_typ : (E : env) (x : atom) (T : qtyp),
      wf_env E
      wf_qtyp E T
      x `notin` dom E
      wf_env (x bind_typ T E)
  | wf_env_qua : (E : env) (X : atom) (Q : qua),
      wf_env E
      wf_qua E Q
      X `notin` dom E
      wf_env (X bind_qua Q E)
.

(* ********************************************************************** *)

Subtyping

The definition of subtyping is straightforward. It uses the binds relation from the MetatheoryEnv library (in the sub_trans_tvar case) and cofinite quantification (in the sub_all case).
Subqualification is simply adapted from standard subtyping rules, and shouldn't appear too surprising as well. Notable exceptions being the two transitivity rules through base lattice elements subqual_eval_elim and subqual_eval_intro, and the lifting rule subqual_lift_elem.

Inductive subqual : env qua qua Prop :=
  | subqual_top : E Q,
      wf_env E
      wf_qua E Q
      subqual E Q (qua_base top)
  | subqual_bot : E Q,
      wf_env E
      wf_qua E Q
      subqual E (qua_base bot) Q
  | subqual_eval_elim : E Q' Q R b,
      subqual E Q Q'
      concretize Q' = Some b
      subqual E (qua_base b) R
      subqual E Q R
  | subqual_eval_intro : E Q' Q R b,
      subqual E Q' R
      concretize Q' = Some b
      subqual E Q (qua_base b)
      subqual E Q R
  | subqual_lift_elem : E b1 b2,
      wf_env E
      ord
      subqual E (qua_base ) (qua_base )
  | subqual_refl_qvar : E X,
      wf_env E
      wf_qua E (qua_fvar X)
      subqual E (qua_fvar X) (qua_fvar X)
  | subqual_trans_qvar : R E Q X,
      binds X (bind_qua R) E
      subqual E R Q
      subqual E (qua_fvar X) Q
  | subqual_join_inl : E R1 R2 Q,
      subqual E Q
      wf_qua E
      subqual E Q (qua_join )
  | subqual_join_inr : E R1 R2 Q,
      wf_qua E
      subqual E Q
      subqual E Q (qua_join )
  | subqual_join_elim : E R1 R2 Q,
      subqual E Q
      subqual E Q
      subqual E (qua_join ) Q
  | subqual_meet_eliml : E R1 R2 Q,
      subqual E Q
      wf_qua E
      subqual E (qua_meet ) Q
  | subqual_meet_elimr : E R1 R2 Q,
      wf_qua E
      subqual E Q
      subqual E (qua_meet ) Q
  | subqual_meet_intro : E R1 R2 Q,
      subqual E Q
      subqual E Q
      subqual E Q (qua_meet )
.

Inductive sub : env typ typ Prop :=
  | sub_top : E S,
      wf_env E
      wf_typ E S
      sub E S typ_top
  | sub_refl_tvar : E X,
      wf_env E
      wf_typ E (typ_fvar X)
      sub E (typ_fvar X) (typ_fvar X)
  | sub_trans_tvar : U E T X,
      binds X (bind_sub U) E
      sub E U T
      sub E (typ_fvar X) T
  | sub_arrow : E S1 S2 T1 T2,
      subqtype E
      subqtype E
      sub E (typ_arrow ) (typ_arrow )
  | sub_all : L E S1 S2 T1 T2,
      sub E
      ( X : atom, X `notin` L
          subqtype (X bind_sub E) (open_tqt X) (open_tqt X))
      sub E (typ_all ) (typ_all )
  | sub_qall : L E S1 S2 T1 T2,
      subqual E
      ( X : atom, X `notin` L
          subqtype (X bind_qua E) (open_qqt X) (open_qqt X))
      sub E (typ_qall ) (typ_qall )
  | sub_sum : E S1 S2 T1 T2,
      subqtype E
      subqtype E
      sub E (typ_sum ) (typ_sum )
  | sub_pair : E S1 S2 T1 T2,
      subqtype E
      subqtype E
      sub E (typ_pair ) (typ_pair )
with subqtype : env qtyp qtyp Prop :=
  | sub_qtyp_qtyp : E Q1 T1 Q2 T2,
      subqual E
      sub E
      subqtype E (qtyp_qtyp ) (qtyp_qtyp )
.

(* ********************************************************************** *)

Typing

The definition of typing is straightforward. It uses the binds relation from the MetatheoryEnv library (in the typing_var case) and cofinite quantification in the cases involving binders (e.g., typing_abs and typing_tabs).

Inductive typing : env exp qtyp Prop :=
  | typing_var : E x T,
      wf_env E
      binds x (bind_typ T) E
      typing E (exp_fvar x) T
  | typing_abs : L E P V e1 T1,
      wf_qua E P
      ( x : atom, x `notin` L
        typing (x bind_typ V E) (open_ee x) )
      typing E (exp_abs P V ) (qtyp_qtyp P (typ_arrow V ))
  | typing_app : Q T1 E e1 e2 T2,
      typing E (qtyp_qtyp Q (typ_arrow ))
      typing E
      typing E (exp_app )
  | typing_tabs : L P E V e1 T1,
      wf_qua E P
      ( X : atom, X `notin` L
        typing (X bind_sub V E) (open_te X) (open_tqt X))
      typing E (exp_tabs P V ) (qtyp_qtyp P (typ_all V ))
  | typing_tapp : Q T1 E e1 T T2,
      typing E (qtyp_qtyp Q (typ_all ))
      sub E T
      typing E (exp_tapp T) (open_tqt T)
  | typing_qabs : L E P Q e1 T1,
      wf_qua E P
      ( X : atom, X `notin` L
        typing (X bind_qua Q E) (open_qe X) (open_qqt X))
      typing E (exp_qabs P Q ) (qtyp_qtyp P (typ_qall Q ))
  | typing_qapp : R Q E e1 Q1 T,
      typing E (qtyp_qtyp R (typ_qall T))
      subqual E Q
      typing E (exp_qapp Q) (open_qqt T Q)
  | typing_sub : S E e T,
      typing E e S
      subqtype E S T
      typing E e T
  | typing_let : L T1 T2 e1 e2 E,
      typing E
      ( x : atom, x `notin` L
        typing (x bind_typ E) (open_ee x) )
      typing E (exp_let )
  | typing_inl : P T1 T2 e1 E,
      wf_qua E P
      typing E
      wf_qtyp E
      typing E (exp_inl P ) (qtyp_qtyp P (typ_sum ))
  | typing_inr : P T1 T2 e1 E,
      wf_qua E P
      typing E
      wf_qtyp E
      typing E (exp_inr P ) (qtyp_qtyp P (typ_sum ))
  | typing_case : L Q T1 T2 T e1 e2 e3 E,
      typing E (qtyp_qtyp Q (typ_sum ))
      ( x : atom, x `notin` L
        typing (x bind_typ E) (open_ee x) T)
      ( x : atom, x `notin` L
        typing (x bind_typ E) (open_ee x) T)
      typing E (exp_case ) T
  | typing_pair : P T1 T2 e1 e2 E,
      wf_qua E P
      typing E
      typing E
      typing E (exp_pair P ) (qtyp_qtyp P (typ_pair ))
  | typing_first : Q T1 T2 e E,
      typing E e (qtyp_qtyp Q (typ_pair ))
      typing E (exp_first e)
  | typing_second : Q T1 T2 e E,
      typing E e (qtyp_qtyp Q (typ_pair ))
      typing E (exp_second e)
  | typing_upqual : E Q T P e,
      typing E e (qtyp_qtyp Q T)
      subqual E Q P
      typing E (exp_upqual P e) (qtyp_qtyp P T)
  | typing_check : E Q T P e,
      typing E e (qtyp_qtyp Q T)
      subqual E Q P
      typing E (exp_check P e) (qtyp_qtyp Q T)
.

(* ********************************************************************** *)

Values

Unlike standard System F-sub, values now carry a qualifier tag.
Inductive value : exp Prop :=
  | value_abs : P T e1,
      expr (exp_abs P T )
      value (exp_abs P T )
  | value_tabs : P T e1,
      expr (exp_tabs P T )
      value (exp_tabs P T )
 | value_qabs : P Q e1,
      expr (exp_qabs P Q )
      value (exp_qabs P Q )
  | value_inl : P e1,
      qual P
      value
      value (exp_inl P )
  | value_inr : P e1,
      qual P
      value
      value (exp_inr P )
  | value_pair : P e1 e2,
      qual P
      value
      value
      value (exp_pair P )
.

(* ********************************************************************** *)

Reduction


Inductive red : exp exp Prop :=
  | red_app_1 : e1 e1' e2,
      expr
      red
      red (exp_app ) (exp_app )
  | red_app_2 : e1 e2 e2',
      value
      red
      red (exp_app ) (exp_app )
  | red_tapp : e1 e1' V,
      type V
      red
      red (exp_tapp V) (exp_tapp V)
  | red_qapp : e1 e1' R,
      qual R
      red
      red (exp_qapp R) (exp_qapp R)
  | red_abs : P T e1 v2,
      expr (exp_abs P T )
      value
      red (exp_app (exp_abs P T ) ) (open_ee )
  | red_tabs : P T1 e1 T2,
      expr (exp_tabs P )
      type
      red (exp_tapp (exp_tabs P ) ) (open_te )
  | red_qabs : P Q1 e1 Q2,
      expr (exp_qabs P )
      qual
      red (exp_qapp (exp_qabs P ) ) (open_qe )
  | red_let_1 : e1 e1' e2,
      red
      body_e
      red (exp_let ) (exp_let )
  | red_let : v1 e2,
      value
      body_e
      red (exp_let ) (open_ee )
  | red_inl_1 : P e1 e1',
      qual P
      red
      red (exp_inl P ) (exp_inl P )
  | red_inr_1 : P e1 e1',
      qual P
      red
      red (exp_inr P ) (exp_inr P )
  | red_case_1 : e1 e1' e2 e3,
      red
      body_e
      body_e
      red (exp_case ) (exp_case )
  | red_case_inl : P v1 e2 e3,
      qual P
      value
      body_e
      body_e
      red (exp_case (exp_inl P ) ) (open_ee )
  | red_case_inr : P v1 e2 e3,
      qual P
      value
      body_e
      body_e
      red (exp_case (exp_inr P ) ) (open_ee )
  | red_pair_1 : P e1 e1' e2,
      qual P
      red
      expr
      red (exp_pair P ) (exp_pair P )
  | red_pair_2 : P v1 e2 e2',
      qual P
      value
      red
      red (exp_pair P ) (exp_pair P )
  | red_pair_first_1 : e1 e1',
      red
      red (exp_first ) (exp_first )
  | red_pair_second_1 : e1 e1',
      red
      red (exp_second ) (exp_second )
  | red_pair_first_2 : P v1 v2,
      qual P
      value
      value
      red (exp_first (exp_pair P ))
  | red_pair_second_2 : P v1 v2,
      qual P
      value
      value
      red (exp_second (exp_pair P ))
The upqual and check rules ensure that concrete qualifiers are compatible before stepping. This ensures that progress/preservation are meaningful -- ill-typed/qualified programs will get stuck.
  | red_upqual_1 : P e1 e1',
      qual P
      red
      red (exp_upqual P ) (exp_upqual P )
  | red_upqual_abs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_abs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_abs Q V )) (exp_abs P V )
  | red_upqual_tabs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_tabs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_tabs Q V )) (exp_tabs P V )
  | red_upqual_qabs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_qabs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_qabs Q V )) (exp_qabs P V )
   | red_upqual_inl : P Q cP cQ e1,
      qual P
      qual Q
      expr (exp_inl Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_inl Q )) (exp_inl P )
   | red_upqual_inr : P Q cP cQ e1,
      qual P
      qual Q
      expr (exp_inr Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_inr Q )) (exp_inr P )
   | red_upqual_pair : P Q cP cQ e1 e2,
      qual P
      qual Q
      expr (exp_pair Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_upqual P (exp_pair Q )) (exp_pair P )
  | red_check_1 : P e1 e1',
      qual P
      red
      red (exp_check P ) (exp_check P )
  | red_check_abs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_abs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_abs Q V )) (exp_abs Q V )
  | red_check_tabs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_tabs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_tabs Q V )) (exp_tabs Q V )
  | red_check_qabs : P Q cP cQ V e1,
      qual P
      qual Q
      expr (exp_qabs Q V )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_qabs Q V )) (exp_qabs Q V )
   | red_check_inl : P Q cP cQ e1,
      qual P
      qual Q
      expr (exp_inl Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_inl Q )) (exp_inl Q )
   | red_check_inr : P Q cP cQ e1,
      qual P
      qual Q
      expr (exp_inr Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_inr Q )) (exp_inr Q )
   | red_check_pair : P Q cP cQ e1 e2,
      qual P
      qual Q
      expr (exp_pair Q )
      concretize P = Some cP
      concretize Q = Some cQ
      cQ cP
      red (exp_check P (exp_pair Q )) (exp_pair Q )
.

(* ********************************************************************** *)

Automation

We declare most constructors as Hints to be used by the auto and eauto tactics. We exclude constructors from the subtyping and typing relations that use cofinite quantification. It is unlikely that eauto will find an instantiation for the finite set L, and in those cases, eauto can take some time to fail. (A priori, this is not obvious. In practice, one adds as hints all constructors and then later removes some constructors when they cause proof search to take too long.)

#[export] Hint Constructors type qtype expr qual wf_qua wf_typ wf_qtyp wf_env value red : core.
#[export] Hint Constructors subqual subqtype: core.
#[export] Hint Resolve sub_top sub_refl_tvar sub_arrow : core.
#[export] Hint Resolve sub_sum sub_pair : core.
#[export] Hint Resolve typing_var typing_app typing_tapp typing_sub typing_qapp : core.
#[export] Hint Resolve typing_inl typing_inr typing_pair typing_first typing_second : core.
#[export] Hint Resolve typing_upqual typing_check : core.