made the fv-definition aggree more with alpha (needed in the support proofs)
authorChristian Urban <urbanc@in.tum.de>
Fri, 03 Sep 2010 22:22:43 +0800
changeset 2464 f4eba60cbd69
parent 2463 217149972715
child 2465 07ffa4e41659
made the fv-definition aggree more with alpha (needed in the support proofs)
Nominal-General/nominal_library.ML
Nominal/Ex/SingleLet.thy
Nominal/nominal_dt_alpha.ML
Nominal/nominal_dt_rawfuns.ML
--- a/Nominal-General/nominal_library.ML	Fri Sep 03 20:53:09 2010 +0800
+++ b/Nominal-General/nominal_library.ML	Fri Sep 03 22:22:43 2010 +0800
@@ -49,6 +49,7 @@
   val mk_append: term * term -> term
   val mk_union: term * term -> term
   val fold_union: term list -> term
+  val fold_append: term list -> term
   val mk_conj: term * term -> term
   val fold_conj: term list -> term
 
@@ -144,6 +145,8 @@
 
 fun mk_diff (@{term "{}::atom set"}, _) = @{term "{}::atom set"}
   | mk_diff (t1, @{term "{}::atom set"}) = t1
+  | mk_diff (@{term "set ([]::atom list)"}, _) = @{term "set ([]::atom list)"}
+  | mk_diff (t1, @{term "set ([]::atom list)"}) = t1
   | mk_diff (t1, t2) = HOLogic.mk_binop @{const_name minus} (t1, t2)
 
 fun mk_append (t1, @{term "[]::atom list"}) = t1
@@ -152,9 +155,12 @@
 
 fun mk_union (t1, @{term "{}::atom set"}) = t1
   | mk_union (@{term "{}::atom set"}, t2) = t2
+  | mk_union (t1, @{term "set ([]::atom list)"}) = t1
+  | mk_union (@{term "set ([]::atom list)"}, t2) = t2
   | mk_union (t1, t2) = HOLogic.mk_binop @{const_name "sup"} (t1, t2)  
  
 fun fold_union trms = fold_rev (curry mk_union) trms @{term "{}::atom set"}
+fun fold_append trms = fold_rev (curry mk_append) trms @{term "[]::atom list"}
 
 fun mk_conj (t1, @{term "True"}) = t1
   | mk_conj (@{term "True"}, t2) = t2
--- a/Nominal/Ex/SingleLet.thy	Fri Sep 03 20:53:09 2010 +0800
+++ b/Nominal/Ex/SingleLet.thy	Fri Sep 03 22:22:43 2010 +0800
@@ -9,17 +9,17 @@
 nominal_datatype single_let: trm  =
   Var "name"
 | App "trm" "trm"
-| Lam x::"name" t::"trm"  bind (set) x in t
-| Let a::"assg" t::"trm"  bind (set) "bn a" in t
+| Lam x::"name" t::"trm"  bind x in t
+| Let a::"assg" t::"trm"  bind "bn a" in t
 | Foo x::"name" y::"name" t::"trm" t1::"trm" t2::"trm" bind (set) x in y t t1 t2
 | Bar x::"name" y::"name" t::"trm" bind y x in t x y
 | Baz x::"name" t1::"trm" t2::"trm" bind x in t1, bind x in t2 
 and assg =
   As "name" x::"name" t::"trm" bind x in t
 binder
-  bn::"assg \<Rightarrow> atom set"
+  bn::"assg \<Rightarrow> atom list"
 where
-  "bn (As x y t) = {atom x}"
+  "bn (As x y t) = [atom x]"
 
 
 thm single_let.distinct
@@ -34,20 +34,6 @@
 thm single_let.supports
 thm single_let.fsupp
 
-lemma supp_abs_sum:
-  fixes a b::"'a::fs"
-  shows "supp (Abs x a) \<union> supp (Abs x b) = supp (Abs x (a, b))"
-  and   "supp (Abs_res x a) \<union> supp (Abs_res x b) = supp (Abs_res x (a, b))"
-  and   "supp (Abs_lst y a) \<union> supp (Abs_lst y b) = supp (Abs_lst y (a, b))"
-  apply (simp_all add: supp_abs supp_Pair)
-  apply blast+
-  done
-
-
-lemma test: 
-  "(\<exists>p. (bs, x) \<approx>gen (op=) f p (cs, y)) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
-sorry
-
 lemma Abs_eq_iff:
   shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
   and   "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (cs, y))"
@@ -57,126 +43,96 @@
 lemma test2:
   assumes "fv_trm t = supp t" 
   shows "\<forall>p. fv_trm (p \<bullet> t) = supp (p \<bullet> t)"
-sorry
-
-lemma yy:
-  "X = Y \<Longrightarrow> finite X = finite Y" by simp
+apply(rule allI)
+apply(rule_tac p="-p" in permute_boolE)
+apply(perm_simp add: single_let.fv_bn_eqvt permute_minus_cancel)
+apply(rule assms)
+done
 
 
 lemma supp_fv:
   "fv_trm t = supp t \<and> fv_assg as = supp as \<and> fv_bn as = {a. infinite {b. \<not>alpha_bn ((a \<rightleftharpoons> b) \<bullet> as) as}}"
 apply(rule single_let.induct)
-apply(simp_all only: single_let.fv_defs)[2]
-apply(simp_all only: supp_def)[2]
-apply(simp_all only: single_let.perm_simps)[2]
-apply(simp_all only: single_let.eq_iff)[2]
-apply(simp_all only: de_Morgan_conj)[2]
-apply(simp_all only: Collect_disj_eq)[2]
-apply(simp_all only: finite_Un)[2]
-apply(simp_all only: de_Morgan_conj)[2]
-apply(simp_all only: Collect_disj_eq)[2]
+apply(simp_all (no_asm_use) only: single_let.fv_defs)[2]
+apply(simp_all (no_asm_use) only: supp_def)[2]
+apply(simp_all (no_asm_use) only: single_let.perm_simps)[2]
+apply(simp_all (no_asm_use) only: single_let.eq_iff)[2]
+apply(simp_all (no_asm_use) only: de_Morgan_conj)[2]
+apply(simp_all (no_asm_use) only: Collect_disj_eq)[2]
+apply(simp_all (no_asm_use) only: finite_Un)[2]
+apply(simp_all (no_asm_use) only: de_Morgan_conj)[2]
+apply(simp_all (no_asm_use) only: Collect_disj_eq)[2]
+apply(simp)
 --" 1 "
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
 apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
 apply(drule test2)
-apply(simp)
+apply(simp only:)
 -- " 2 "
 apply(erule conjE)+
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
-apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: Abs_eq_iff)
-apply(simp only: finite_Un)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp add: alphas)
-apply(drule test2)
-apply(simp)
--- " 3 "
-apply(simp only: single_let.fv_defs)
-apply(simp only: supp_Pair[symmetric])
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
 apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas)
-apply(simp add: supp_Pair)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(drule test2)
+apply(simp only:)
+-- " 3 "
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(1)[symmetric])
+apply(simp (no_asm) only: supp_def)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
 apply(drule test2)+
-apply(simp)
-apply(simp add: prod_alpha_def)
-apply(simp add: Un_assoc)
-apply(rule Collect_cong)
-apply(rule arg_cong)
-back
-apply(rule yy)
-apply(rule Collect_cong)
-apply(auto)[1]
+apply(simp only: supp_Pair Un_assoc conj_assoc)
 -- " Bar "
-apply(simp only: single_let.fv_defs)
-apply(simp only: supp_Pair[symmetric])
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
 apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas prod_alpha_def)
-apply(drule test2)
-apply(simp add: supp_Pair)
-apply(subst atom_eqvt)
-apply(simp)
-apply(simp add: Un_assoc)
-apply(rule Collect_cong)
-apply(rule arg_cong)
-back
-apply(rule yy)
-apply(rule Collect_cong)
--- "last"
-prefer 3
-apply(rule conjI)
-apply(simp only: single_let.fv_defs)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp add: supp_abs(1)[symmetric])
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- " Baz "
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
 apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: finite_Un)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp add: alphas prod_alpha_def)
-apply(drule test2)
-apply(simp add: supp_Pair)
-apply(simp only: permute_set_eq)
-apply(simp)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: single_let.eq_iff)
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- "last"
+apply(rule conjI)
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
 apply(simp (no_asm) only: supp_def)
-apply(perm_simp)
-oops
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- "other case"
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
+apply(simp (no_asm) only: supp_def)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)?
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+done
 
 
 
--- a/Nominal/nominal_dt_alpha.ML	Fri Sep 03 20:53:09 2010 +0800
+++ b/Nominal/nominal_dt_alpha.ML	Fri Sep 03 22:22:43 2010 +0800
@@ -64,7 +64,7 @@
 end
 
 (* generates the compound binder terms *)
-fun mk_binders lthy bmode args bodies = 
+fun mk_binders lthy bmode args binders = 
 let  
   fun bind_set lthy args (NONE, i) = setify lthy (nth args i)
     | bind_set _ args (SOME bn, i) = bn $ (nth args i)
@@ -77,7 +77,7 @@
     | Set => (mk_union,  bind_set)
     | Res => (mk_union,  bind_set)
 in
-  bodies
+  binders
   |> map (bind_fn lthy args)
   |> foldl1 combine_fn 
 end
--- a/Nominal/nominal_dt_rawfuns.ML	Fri Sep 03 20:53:09 2010 +0800
+++ b/Nominal/nominal_dt_rawfuns.ML	Fri Sep 03 22:22:43 2010 +0800
@@ -151,7 +151,7 @@
 
 (** functions that construct the equations for fv and fv_bn **)
 
-fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (_, binders, bodies)) =
+fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (bmode, binders, bodies)) =
 let
   fun mk_fv_body fv_map args i = 
   let
@@ -163,21 +163,33 @@
     | SOME fv => fv $ arg
   end  
 
-  fun mk_fv_binder lthy fv_bn_map args (bn_option, i) = 
+  fun mk_fv_binder lthy fv_bn_map args binders = 
   let
-    val arg = nth args i
+    fun bind_set lthy args (NONE, i) = (setify lthy (nth args i), @{term "{}::atom set"})
+      | bind_set _ args (SOME bn, i) = (bn $ (nth args i), 
+          if  member (op=) bodies i then @{term "{}::atom set"}  
+          else lookup fv_bn_map bn $ (nth args i))
+    fun bind_lst lthy args (NONE, i) = (listify lthy (nth args i), @{term "[]::atom list"})
+      | bind_lst _ args (SOME bn, i) = (bn $ (nth args i),
+          if  member (op=) bodies i then @{term "[]::atom list"}  
+          else lookup fv_bn_map bn $ (nth args i)) 
+  
+    val (combine_fn, bind_fn) =
+      case bmode of
+        Lst => (fold_append, bind_lst) 
+      | Set => (fold_union, bind_set)
+      | Res => (fold_union, bind_set)
   in
-    case bn_option of
-      NONE => (setify lthy arg, @{term "{}::atom set"})
-    | SOME bn => (to_set (bn $ arg), 
-       if  member (op=) bodies i then @{term "{}::atom set"}  
-       else lookup fv_bn_map bn $ arg)
+    binders
+    |> map (bind_fn lthy args)
+    |> split_list
+    |> pairself combine_fn
   end  
 
   val t1 = map (mk_fv_body fv_map args) bodies
-  val (t2, t3) = split_list (map (mk_fv_binder lthy fv_bn_map args) binders)
+  val (t2, t3) = mk_fv_binder lthy fv_bn_map args binders
 in 
-  fold_union (mk_diff (fold_union t1, fold_union t2)::t3)
+  mk_union (mk_diff (fold_union t1, to_set t2), to_set t3)
 end
 
 (* in case of fv_bn we have to treat the case special, where an