How to define Sup for an inductive datatype?

2019-08-22 22:17发布

问题:

Here is a simple type system with following types: any, void, integer, real, set.

datatype ty =
  AType
| VType
| IType
| RType
| SType ty

Types form a semilattice with supremum:

notation sup (infixl "⊔" 65)

instantiation ty :: semilattice_sup
begin

inductive less_ty where
  "τ ≠ VType ⟹ VType < τ"
| "τ ≠ AType ⟹ τ < AType"
| "IType < RType"
| "τ < σ ⟹ SType τ < SType σ"

inductive_cases VType_less [elim!]: "VType < τ"
inductive_cases less_AType [elim!]: "τ < AType"
inductive_cases IType_less_RType [elim!]: "IType < RType"
inductive_cases SType_less_SType [elim!]: "SType τ < SType σ"
inductive_cases less_SType [elim!]: "τ < SType σ"
inductive_cases SType_less [elim!]: "SType τ < σ"

fun less_ty_fun where
  "less_ty_fun VType τ = (τ ≠ VType)"
| "less_ty_fun τ VType = False"
| "less_ty_fun AType τ = False"
| "less_ty_fun τ AType = (τ ≠ AType)"
| "less_ty_fun IType RType = True"
| "less_ty_fun (SType τ) (SType σ) = (τ < σ)"
| "less_ty_fun _ _ = False"

lemma less_ty_code [code]:
  "τ < σ = less_ty_fun τ σ"
  for τ σ :: ty
  apply (rule iffI)
  apply (induct rule: less_ty.induct)
  apply simp
  using less_ty_fun.simps(10) ty.exhaust apply force
  apply simp
  apply simp
  apply (erule less_ty_fun.elims; simp add: less_ty.intros)
  done

definition "τ ≤ σ ≡ τ = σ ∨ τ < σ" for τ σ :: ty

fun sup_ty where
  "VType ⊔ τ = τ"
| "τ ⊔ VType = τ"
| "AType ⊔ τ = AType"
| "τ ⊔ AType = AType"
| "IType ⊔ IType = IType"
| "IType ⊔ RType = RType"
| "RType ⊔ IType = RType"
| "RType ⊔ RType = RType"
| "SType τ ⊔ SType σ = SType (τ ⊔ σ)"
| "SType _ ⊔ _ = AType"
| "_ ⊔ SType _ = AType"

lemma less_le_not_le_ty:
  "τ < σ ⟷ τ ≤ σ ∧ ¬ σ ≤ τ" for τ σ :: ty
  apply (rule iffI; auto simp add: less_eq_ty_def)
  apply (induct τ arbitrary: σ; auto)
  using less_ty.cases apply fastforce
  using less_ty.cases apply fastforce
  using less_ty.cases apply fastforce
  apply (induct rule: less_ty.induct)
  using less_ty.cases apply blast+
  done

lemma order_refl_ty [iff]: "τ ≤ τ" for τ :: ty
  by (simp add: less_eq_ty_def)

lemma order_trans_ty:
  "τ ≤ σ ⟹ σ ≤ ρ ⟹ τ ≤ ρ" for τ σ ρ :: ty
  apply (auto simp add: less_eq_ty_def)
  apply (erule notE)
  apply (induct τ arbitrary: σ ρ)
  using less_ty.cases apply blast
  apply (metis less_le_not_le_ty less_ty.intros(1))
  apply (metis less_ty.simps ty.distinct(1) ty.distinct(19) ty.distinct(3) ty.distinct(7) ty.simps(13) ty.simps(19))
  apply (metis less_ty.simps ty.distinct(1) ty.distinct(19) ty.distinct(3) ty.distinct(7) ty.simps(13))
  apply (metis less_ty.simps ty.distinct(1) ty.distinct(17) ty.distinct(3) ty.distinct(7) ty.inject ty.simps(15))
  done

lemma antisym_ty:
  "τ ≤ σ ⟹ σ ≤ τ ⟹ τ = σ" for τ σ :: ty
  by (simp add: less_eq_ty_def; auto simp add: less_le_not_le_ty)

lemma sup_ge1_ty_IType:
  "IType = IType ⊔ σ ∨ IType < IType ⊔ σ"
  by (cases σ; simp add: less_ty.intros)

lemma sup_ge1_ty_RType:
  "RType = RType ⊔ σ ∨ RType < RType ⊔ σ"
  by (cases σ; simp add: less_ty.intros)

lemma sup_ge1_ty_SType:
  "(⋀σ. τ = τ ⊔ σ ∨ τ < τ ⊔ σ) ⟹
   SType τ = SType τ ⊔ σ ∨ SType τ < SType τ ⊔ σ"
  apply (cases σ; simp add: less_ty.intros)
  using less_ty.intros(4) by auto

lemma sup_ge1_ty: "τ ≤ τ ⊔ σ" for τ σ :: ty
  apply (induct τ arbitrary: σ; simp add: less_eq_ty_def)
  apply (metis sup_ty.simps(2) sup_ty.simps(6) sup_ty.simps(7) sup_ty.simps(8) sup_ty.simps(9) ty.exhaust)
  apply (metis less_ty.intros(1))
  using sup_ge1_ty_IType apply auto[1]
  using sup_ge1_ty_RType apply auto[1]
  using sup_ge1_ty_SType apply auto[1]
  done

lemma sup_commut_ty_AType: "AType ⊔ σ = σ ⊔ AType" by (cases σ; simp)
lemma sup_commut_ty_VType: "VType ⊔ σ = σ ⊔ VType" by (cases σ; simp)
lemma sup_commut_ty_IType: "IType ⊔ σ = σ ⊔ IType" by (cases σ; simp)
lemma sup_commut_ty_RType: "RType ⊔ σ = σ ⊔ RType" by (cases σ; simp)
lemma sup_commut_ty_SType: "(⋀σ. τ ⊔ σ = σ ⊔ τ) ⟹ SType τ ⊔ σ = σ ⊔ SType τ" by (cases σ; simp)

lemma sup_commut_ty:
  "τ ⊔ σ = σ ⊔ τ"
  for τ σ :: ty
  apply (induct τ arbitrary: σ)
  using sup_commut_ty_AType apply auto[1]
  using sup_commut_ty_VType apply auto[1]
  using sup_commut_ty_IType apply auto[1]
  using sup_commut_ty_RType apply auto[1]
  using sup_commut_ty_SType apply auto[1]
  done

lemma sup_ty_idem: "τ ⊔ τ = τ" for τ :: ty by (induct τ; simp)

lemma sup_ty_strict_order:
  "σ < τ ⟹ τ ⊔ σ = τ"
  for τ σ :: ty
  apply (induct rule: less_ty.induct)
  using sup_commut_ty sup_ty.simps(1) apply auto[1]
  apply (meson antisym_ty less_eq_ty_def less_ty.simps sup_ge1_ty)
  apply simp
  apply simp
  done

lemma sup_less_eq_set:
  "(⋀τ σ. τ < ρ ⟹ σ < ρ ⟹ τ ⊔ σ ≤ ρ) ⟹
   τ < SType ρ ⟹
   σ < SType ρ ⟹
   τ ⊔ σ ≤ SType ρ"
  apply (cases τ; cases σ; auto)
  apply (simp add: less_eq_ty_def less_ty.intros(1) sup_ty_idem)
  apply (simp add: less_eq_ty_def less_ty.intros(4))
  apply (simp add: less_eq_ty_def less_ty.intros(4))
  apply (metis (mono_tags) less_eq_ty_def less_ty.intros(4))
  done

lemma sup_ty_strict_order2_RType:
  "τ < RType ⟹ σ < RType ⟹ τ ⊔ σ ≤ RType"
  apply (cases τ; auto)
  apply (simp add: less_ty_code)
  apply (simp add: less_eq_ty_def)
  apply (metis less_eq_ty_def less_ty.simps sup_ty.simps(13) sup_ty.simps(3) ty.distinct(19) ty.distinct(5))
  apply (simp add: less_ty_code)
  done

lemma sup_ty_strict_order2:
  "τ < ρ ⟹ σ < ρ ⟹ τ ⊔ σ ≤ ρ" for τ σ ρ :: ty
  apply (induct ρ arbitrary: τ σ)
  using less_eq_ty_def less_ty.intros(2) apply blast
  using less_ty.cases apply blast
  using less_le_not_le_ty less_ty.cases apply fastforce
  apply (simp add: sup_ty_strict_order2_RType)
  apply (simp add: sup_less_eq_set)
  done

lemma sup_least_ty:
  "τ ≤ ρ ⟹ σ ≤ ρ ⟹ τ ⊔ σ ≤ ρ" for τ σ ρ :: ty
  apply (simp add: less_eq_ty_def)
  apply (elim disjE)
  using sup_ty_idem apply auto[1]
  apply (simp add: sup_ty_strict_order)
  apply (simp add: sup_ty_strict_order sup_commut_ty)
  using less_eq_ty_def sup_ty_strict_order2 apply auto
  done

instance
  apply (standard)
  apply (simp add: less_le_not_le_ty)
  apply simp
  using order_trans_ty apply blast
  apply (simp add: antisym_ty)
  apply (simp add: sup_ge1_ty)
  apply (simp add: sup_commut_ty sup_ge1_ty)
  apply (simp add: sup_least_ty)
  done

end

Everything works as expected:

value "IType ⊔ RType"
value "SType IType ⊔ SType RType"
value "SType IType ⊔ SType (SType RType)"
value "SType (SType IType) ⊔ SType (SType RType)"

I try to define a Sup function for types:

interpretation ty_abel_semigroup: abel_semigroup "sup :: ty ⇒ ty ⇒ ty" ..

interpretation ty_comm_monoid_set: comm_monoid_set "sup :: ty ⇒ ty ⇒ ty" VType
  apply (standard)
  using sup_commut_ty sup_ty.simps(1) by auto

instantiation ty :: Sup
begin
definition Sup_ty where "Sup_ty ≡ ty_comm_monoid_set.F id"
instance ..
end

The problem is that the following expression can't be evaluated:

value "Sup {IType, RType}"

I think this is caused by the fact that ty data type is infinite and so a set of ty is infinite too.

I tried to describe UNIV :: ty set as follows and prove it's finiteness:

lemma UNIV_ty: "UNIV = {VType, AType, IType, RType} ∪ (SType ` UNIV)"
  apply (auto intro: ty.induct)
  by (metis range_eqI ty.exhaust)

instance ty :: finite
  apply (standard)

But I'm stuck. I'm not sure whether it's finite or not. Finite lists are defined as inductive datatypes too. Are all inductive datatypes finite?

UPDATE:

I proved that ty is countable. But I have no idea whether it can help to define Sup for ty...

fun ty_to_nat :: "ty ⇒ nat" where
  "ty_to_nat VType = 0"
| "ty_to_nat AType = 1"
| "ty_to_nat IType = 2"
| "ty_to_nat RType = 3"
| "ty_to_nat (SType t) = 4 + ty_to_nat t"

lemma ty_to_nat_inj_AType: "ty_to_nat AType = ty_to_nat y ⟹ AType = y"
  by (induct y; simp)
lemma ty_to_nat_inj_VType: "ty_to_nat VType = ty_to_nat y ⟹ VType = y"
  by (induct y; simp)
lemma ty_to_nat_inj_IType: "ty_to_nat IType = ty_to_nat y ⟹ IType = y"
  by (induct y; simp)
lemma ty_to_nat_inj_RType: "ty_to_nat RType = ty_to_nat y ⟹ RType = y"
  by (induct y; simp)
lemma ty_to_nat_inj_SType:
  "(⋀y. ty_to_nat x = ty_to_nat y ⟹ x = y) ⟹
   ty_to_nat (SType x) = ty_to_nat y ⟹ SType x = y"
  by (induct y; simp)

lemma ty_to_nat_inj:
  "ty_to_nat x = ty_to_nat y ⟹ x = y"
  apply (induct x arbitrary: y)
  using ty_to_nat_inj_AType apply auto[1]
  using ty_to_nat_inj_VType apply auto[1]
  using ty_to_nat_inj_IType apply auto[1]
  using ty_to_nat_inj_RType apply auto[1]
  using ty_to_nat_inj_SType apply auto[1]
  done

instance ty :: countable
  apply (standard)
  apply (rule exI[of _ "ty_to_nat"])
  apply (simp add: injI ty_to_nat_inj)
  done

回答1:

It is not possible to compute suprema of infinite sets unless you know something about the structure of the infinite set to which you apply the Sup operator. So there is no general-purpose solution.

You can of course check using non-executable membership tests that are tailored to your order definition. You may want to look at the DFA theories of Jinja and JinjaThreads (available in the AFP) where least upper bounds are defined and computed for a Java-like class hierarchy.

For execution, if you are only interested in suprema over finite sets, you can derive a special code equaiton which pattern-matches on the set code constructor. For example, you could prove a code equation similar to Sup_set_fold in theory List. This particular theorem requires a least element, too, and I haven't studied your example enough to see whethere there is such.



标签: isabelle