Theory Unary_PCF
section ‹Towards Decidability of Behavioral Equivalence for Unary PCF›
theory Unary_PCF
imports
"HOL-Library.FSet"
"HOL-Library.Countable_Set_Type"
"HOL-Library.Nat_Bijection"
Hereditary_Multiset
"List-Index.List_Index"
begin
subsection ‹Preliminaries›
lemma prod_UNIV: "UNIV = UNIV × UNIV"
by auto
lemma infinite_cartesian_productI1: "infinite A ⟹ B ≠ {} ⟹ infinite (A × B)"
by (auto dest!: finite_cartesian_productD1)
subsection ‹Types›
datatype type = ℬ ("ℬ") | Fun type type (infixr "→" 65)
definition mk_fun (infixr "→→" 65) where
"Ts →→ T = fold (→) (rev Ts) T"
primrec dest_fun where
"dest_fun ℬ = []"
| "dest_fun (T → U) = T # dest_fun U"
definition arity where
"arity T = length (dest_fun T)"
lemma mk_fun_dest_fun[simp]: "dest_fun T →→ ℬ = T"
by (induct T) (auto simp: mk_fun_def)
lemma dest_fun_mk_fun[simp]: "dest_fun (Ts →→ T) = Ts @ dest_fun T"
by (induct Ts) (auto simp: mk_fun_def)
primrec δ where
"δ ℬ = HMSet {#}"
| "δ (T → U) = HMSet (add_mset (δ T) (hmsetmset (δ U)))"
lemma δ_mk_fun: "δ (Ts →→ T) = HMSet (hmsetmset (δ T) + mset (map δ Ts))"
by (induct Ts) (auto simp: mk_fun_def)
lemma type_induct [case_names Fun]:
assumes
"(⋀T. (⋀T1 T2. T = T1 → T2 ⟹ P T1) ⟹
(⋀T1 T2. T = T1 → T2 ⟹ P T2) ⟹ P T)"
shows "P T"
proof (induct T)
case ℬ
show ?case by (rule assms) simp_all
next
case Fun
show ?case by (rule assms) (insert Fun, simp_all)
qed
subsection ‹Terms›
type_synonym name = string
type_synonym idx = nat