Module Qcert.TypeSystem.TBrandModel
Require Import String.
Require Import List.
Require Import ListSet.
Require Import Sumbool.
Require Import Arith.
Require Import Bool.
Require Import Morphisms.
Require Import Basics.
Require Import EquivDec.
Require Import Eqdep_dec.
Require Import Utils.
Require Import ForeignType.
Require Import RType.
Require Export BrandRelation.
Require Export TBrandContext.
Require Import RSubtype.
Require Import RConsistentSubtype.
Section TBrandModel.
Definition sub_brand_context {
br:
brand_relation} {
ftype:
foreign_type} (
m1 m2:
brand_context) :=
subtype (
brand_context_Rec m1) (
brand_context_Rec m2).
Section brand_model_types.
Context {
ftype:
foreign_type} {
br:
brand_relation} {
m:
brand_context}.
Definition brand_model_domain_t
:=
forall a,
In a (
domain brand_relation_brands) ->
In a (
domain brand_context_types).
Definition brand_model_subtype_weak_t :=
forall a b τ
a,
In (
a,
b)
brand_relation_brands ->
In (
a,τ
a)
brand_context_types ->
exists (τ
b:
rtype),
In (
b,τ
b)
brand_context_types /\
subtype τ
a τ
b.
Definition brand_model_subtype_t :=
forall a b τ
a,
In (
a,
b)
brand_relation_brands ->
In (
a,τ
a)
brand_context_types ->
{τ
b:
rtype |
In (
b,τ
b)
brand_context_types /\
subtype τ
a τ
b}.
Section brand_model_types_dec.
Lemma brand_model_domain_dec :
{
brand_model_domain_t }
+ {~
brand_model_domain_t}.
Proof.
Lemma brand_model_subtype_weak_dec :
{
brand_model_subtype_weak_t}
+ {~
brand_model_subtype_weak_t}.
Proof.
End brand_model_types_dec.
End brand_model_types.
Context {
ftype:
foreign_type}.
Class brand_model :=
mkBrand_model {
brand_model_relation :>
brand_relation;
brand_model_context :>
brand_context;
brand_model_domain_b :
holds (
brand_model_domain_dec);
brand_model_subtype_weak_b :
holds (
brand_model_subtype_weak_dec)
}.
Section brand_model_prop.
Context {
m:
brand_model}.
Lemma brand_model_domain :
brand_model_domain_t.
Proof.
Lemma brand_model_subtype_weak :
brand_model_subtype_weak_t.
Proof.
Lemma brand_model_subtype :
brand_model_subtype_t.
Proof.
End brand_model_prop.
Lemma brand_model_ext a b pf1 pf2 pf1'
pf2' :
mkBrand_model a b pf1 pf2 =
mkBrand_model a b pf1'
pf2'.
Proof.
Definition make_brand_model (
b:
brand_relation) (
m:
brand_context)
(
pf:
(
is_true (
brand_model_domain_dec)
&&
is_true (
brand_model_subtype_weak_dec)
) =
true) :
brand_model.
Proof.
unfold is_true in *.
match_case_in pf;
intros ?
eqq.
-
case_eq (
brand_model_subtype_weak_dec);
intros ?
eqq2.
+
apply (@
mkBrand_model b m);
unfold holds,
is_true;
try rewrite eqq;
try rewrite eqq2;
trivial.
+
rewrite eqq,
eqq2 in pf.
simpl in pf.
discriminate.
-
rewrite eqq in pf.
simpl in pf.
discriminate.
Defined.
Definition make_brand_model_fails (
b:
brand_relation) (
m:
brand_context) :
option brand_model.
Proof.
Program Definition empty_brand_model :=
make_brand_model (
mkBrand_relation nil _ _) (
mkBrand_context nil _).
Next Obligation.
Next Obligation.
unfold holds.
auto.
Defined.
End TBrandModel.