Nominal/Ex/SingleLet.thy
changeset 2448 b9d9c4540265
parent 2436 3885dc2669f9
child 2449 6b51117b8955
equal deleted inserted replaced
2447:76be909eaf04 2448:b9d9c4540265
    19 binder
    19 binder
    20   bn::"assg \<Rightarrow> atom set"
    20   bn::"assg \<Rightarrow> atom set"
    21 where
    21 where
    22   "bn (As x y t) = {atom x}"
    22   "bn (As x y t) = {atom x}"
    23 
    23 
       
    24 thm Ball_def Bex_def mem_def
       
    25 
    24 thm single_let.distinct
    26 thm single_let.distinct
    25 thm single_let.induct
    27 thm single_let.induct
    26 thm single_let.exhaust
    28 thm single_let.exhaust
    27 thm single_let.fv_defs
    29 thm single_let.fv_defs
    28 thm single_let.bn_defs
    30 thm single_let.bn_defs
    29 thm single_let.perm_simps
    31 thm single_let.perm_simps
    30 thm single_let.eq_iff
    32 thm single_let.eq_iff
    31 thm single_let.fv_bn_eqvt
    33 thm single_let.fv_bn_eqvt
    32 thm single_let.size_eqvt
    34 thm single_let.size_eqvt
       
    35 thm single_let.supports
       
    36 
       
    37 lemma test:
       
    38   "finite (supp (t::trm)) \<and> finite (supp (a::assg))"
       
    39 apply(rule single_let.induct)
       
    40 apply(rule supports_finite)
       
    41 apply(rule single_let.supports)
       
    42 apply(simp only: finite_supp supp_Pair finite_Un)
       
    43 apply(rule supports_finite)
       
    44 apply(rule single_let.supports)
       
    45 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    46 apply(rule supports_finite)
       
    47 apply(rule single_let.supports)
       
    48 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    49 apply(rule supports_finite)
       
    50 apply(rule single_let.supports)
       
    51 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    52 apply(rule supports_finite)
       
    53 apply(rule single_let.supports)
       
    54 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    55 apply(rule supports_finite)
       
    56 apply(rule single_let.supports)
       
    57 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    58 apply(rule supports_finite)
       
    59 apply(rule single_let.supports)
       
    60 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    61 apply(rule supports_finite)
       
    62 apply(rule single_let.supports)
       
    63 apply(simp only: finite_supp supp_Pair finite_Un, simp)
       
    64 done
       
    65 
       
    66 instantiation trm and assg :: fs
       
    67 begin
       
    68 
       
    69 instance
       
    70 apply(default)
       
    71 apply(simp_all add: test)
       
    72 done
       
    73 
       
    74 end
    33 
    75 
    34 
    76 
    35 (*
       
    36 
    77 
    37 
    78 
       
    79 lemma test: 
       
    80   "(\<exists>p. (bs, x) \<approx>lst (op=) f p (cs, y)) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>lst (op=) supp p (cs, y))"
       
    81 oops
       
    82 
       
    83 lemma Abs_eq_iff:
       
    84   shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
       
    85   and   "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (cs, y))"
       
    86   and   "Abs_lst bsl x = Abs_lst csl y \<longleftrightarrow> (\<exists>p. (bsl, x) \<approx>lst (op =) supp p (csl, y))"
       
    87   by (lifting alphas_abs)
       
    88 
       
    89 (*
    38 lemma supp_fv:
    90 lemma supp_fv:
    39   "supp t = fv_trm t"
    91   "supp t = fv_trm t \<and> supp b = fv_bn b"
    40   "supp b = fv_bn b"
    92 apply(rule single_let.induct)
    41 apply(induct t and b rule: i1)
    93 apply(simp_all only: single_let.fv_defs)[2]
    42 apply(simp_all add: f1)
    94 apply(simp_all only: supp_def)[2]
    43 apply(simp_all add: supp_def)
    95 apply(simp_all only: single_let.perm_simps)[2]
    44 apply(simp_all add: b1)
    96 apply(simp_all only: single_let.eq_iff)[2]
       
    97 apply(simp_all only: de_Morgan_conj)[2]
       
    98 apply(simp_all only: Collect_disj_eq)[2]
       
    99 apply(simp_all only: finite_Un)[2]
       
   100 apply(simp_all only: de_Morgan_conj)[2]
       
   101 apply(simp_all only: Collect_disj_eq)[2]
       
   102 apply(subgoal_tac "supp (Lam name trm) = supp (Abs_lst [atom name] trm)")
       
   103 apply(simp only: single_let.fv_defs)
       
   104 apply(simp only: supp_abs)
       
   105 apply(simp (no_asm) only: supp_def)
       
   106 apply(simp only: single_let.perm_simps)
       
   107 apply(simp only: single_let.eq_iff)
       
   108 apply(subst test)
       
   109 apply(simp only: Abs_eq_iff[symmetric])
       
   110 apply(simp only: alphas_abs[symmetric])
       
   111 apply(simp only: eqvts)
       
   112 thm Abs_eq_iff
       
   113 apply(simp only: alphas)
    45 sorry
   114 sorry
       
   115 *)
       
   116 (*
    46 
   117 
    47 consts perm_bn_trm :: "perm \<Rightarrow> trm \<Rightarrow> trm"
   118 consts perm_bn_trm :: "perm \<Rightarrow> trm \<Rightarrow> trm"
    48 consts perm_bn_assg :: "perm \<Rightarrow> assg \<Rightarrow> assg"
   119 consts perm_bn_assg :: "perm \<Rightarrow> assg \<Rightarrow> assg"
    49 
   120 
    50 lemma y:
   121 lemma y: