used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on
theory Terms
imports Nominal "../QuotMain" "../QuotList"
begin
atom_decl name
text {*** lets with binding patterns ***}
datatype trm1 =
Vr1 "name"
| Ap1 "trm1" "trm1"
| Lm1 "name" "trm1" --"name is bound in trm1"
| Lt1 "bp" "trm1" "trm1" --"all variables in bp are bound in the 2nd trm1"
and bp =
BUnit
| BVr "name"
| BPr "bp" "bp"
(* to be given by the user *)
fun
bv1
where
"bv1 (BUnit) = {}"
| "bv1 (BVr x) = {x}"
| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"
(* needs to be calculated by the package *)
fun
fv_trm1 and fv_bp
where
"fv_trm1 (Vr1 x) = {x}"
| "fv_trm1 (Ap1 t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2)"
| "fv_trm1 (Lm1 x t) = (fv_trm1 t) - {x}"
| "fv_trm1 (Lt1 bp t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2 - bv1 bp)"
| "fv_bp (BUnit) = {}"
| "fv_bp (BVr x) = {x}"
| "fv_bp (BPr b1 b2) = (fv_bp b1) \<union> (fv_bp b2)"
text {*** lets with single assignments ***}
datatype trm2 =
Vr2 "name"
| Ap2 "trm2" "trm2"
| Lm2 "name" "trm2"
| Lt2 "assign" "trm2"
and assign =
As "name" "trm2"
(* to be given by the user *)
fun
bv2
where
"bv2 (As x t) = {x}"
(* needs to be calculated by the package *)
fun
fv_trm2 and fv_assign
where
"fv_trm2 (Vr2 x) = {x}"
| "fv_trm2 (Ap2 t1 t2) = (fv_trm2 t1) \<union> (fv_trm2 t2)"
| "fv_trm2 (Lm2 x t) = (fv_trm2 t) - {x}"
| "fv_trm2 (Lt2 as t) = (fv_trm2 t - bv2 as) \<union> (fv_assign as)"
| "fv_assign (As x t) = (fv_trm2 t)"
text {*** lets with many assignments ***}
datatype trm3 =
Vr3 "name"
| Ap3 "trm3" "trm3"
| Lm3 "name" "trm3"
| Lt3 "assigns" "trm3"
and assigns =
ANil
| ACons "name" "trm3" "assigns"
(* to be given by the user *)
fun
bv3
where
"bv3 ANil = {}"
| "bv3 (ACons x t as) = {x} \<union> (bv3 as)"
fun
fv_trm3 and fv_assigns
where
"fv_trm3 (Vr3 x) = {x}"
| "fv_trm3 (Ap3 t1 t2) = (fv_trm3 t1) \<union> (fv_trm3 t2)"
| "fv_trm3 (Lm3 x t) = (fv_trm3 t) - {x}"
| "fv_trm3 (Lt3 as t) = (fv_trm3 t - bv3 as) \<union> (fv_assigns as)"
| "fv_assigns (ANil) = {}"
| "fv_assigns (ACons x t as) = (fv_trm3 t) \<union> (fv_assigns as)"
end