Nominal/Ex/SingleLet.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 22 Dec 2010 09:13:25 +0000
changeset 2617 e44551d067e6
parent 2616 dd7490fdd998
child 2622 e6e6a3da81aa
permissions -rw-r--r--
properly exported strong exhaust theorem; cleaned up some examples
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1911
60b5c61d3de2 renamed Ex1.thy to SingleLet.thy
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
     1
theory SingleLet
2474
6e37bfb62474 generated inducts rule by Project_Rule.projections
Christian Urban <urbanc@in.tum.de>
parents: 2473
diff changeset
     2
imports "../Nominal2" 
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
2475
486d4647bb37 supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
parents: 2474
diff changeset
     7
nominal_datatype single_let: trm =
1911
60b5c61d3de2 renamed Ex1.thy to SingleLet.thy
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
     8
  Var "name"
60b5c61d3de2 renamed Ex1.thy to SingleLet.thy
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
     9
| App "trm" "trm"
2464
f4eba60cbd69 made the fv-definition aggree more with alpha (needed in the support proofs)
Christian Urban <urbanc@in.tum.de>
parents: 2461
diff changeset
    10
| Lam x::"name" t::"trm"  bind x in t
f4eba60cbd69 made the fv-definition aggree more with alpha (needed in the support proofs)
Christian Urban <urbanc@in.tum.de>
parents: 2461
diff changeset
    11
| Let a::"assg" t::"trm"  bind "bn a" in t
2424
621ebd8b13c4 changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
Christian Urban <urbanc@in.tum.de>
parents: 2410
diff changeset
    12
| Foo x::"name" y::"name" t::"trm" t1::"trm" t2::"trm" bind (set) x in y t t1 t2
2296
45a69c9cc4cc alpha works now
Christian Urban <urbanc@in.tum.de>
parents: 2295
diff changeset
    13
| Bar x::"name" y::"name" t::"trm" bind y x in t x y
2493
2e174807c891 added postprocessed fresh-lemmas for constructors
Christian Urban <urbanc@in.tum.de>
parents: 2492
diff changeset
    14
| Baz x::"name" t1::"trm" t2::"trm" bind (set) x in t1, bind (res) x in t2 
1911
60b5c61d3de2 renamed Ex1.thy to SingleLet.thy
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    15
and assg =
2320
d835a2771608 prove that alpha implies alpha_bn (needed for rsp proofs)
Christian Urban <urbanc@in.tum.de>
parents: 2318
diff changeset
    16
  As "name" x::"name" t::"trm" bind x in t
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
binder
2464
f4eba60cbd69 made the fv-definition aggree more with alpha (needed in the support proofs)
Christian Urban <urbanc@in.tum.de>
parents: 2461
diff changeset
    18
  bn::"assg \<Rightarrow> atom list"
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
where
2464
f4eba60cbd69 made the fv-definition aggree more with alpha (needed in the support proofs)
Christian Urban <urbanc@in.tum.de>
parents: 2461
diff changeset
    20
  "bn (As x y t) = [atom x]"
1911
60b5c61d3de2 renamed Ex1.thy to SingleLet.thy
Christian Urban <urbanc@in.tum.de>
parents: 1773
diff changeset
    21
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    22
thm single_let.distinct
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    23
thm single_let.induct
2474
6e37bfb62474 generated inducts rule by Project_Rule.projections
Christian Urban <urbanc@in.tum.de>
parents: 2473
diff changeset
    24
thm single_let.inducts
2509
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    25
thm single_let.exhaust[no_vars]
2436
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    26
thm single_let.fv_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    27
thm single_let.bn_defs
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    28
thm single_let.perm_simps
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    29
thm single_let.eq_iff
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    30
thm single_let.fv_bn_eqvt
3885dc2669f9 cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents: 2434
diff changeset
    31
thm single_let.size_eqvt
2448
b9d9c4540265 proved supports lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2436
diff changeset
    32
thm single_let.supports
2450
217ef3e4282e added proofs for fsupp properties
Christian Urban <urbanc@in.tum.de>
parents: 2449
diff changeset
    33
thm single_let.fsupp
2475
486d4647bb37 supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
Christian Urban <urbanc@in.tum.de>
parents: 2474
diff changeset
    34
thm single_let.supp
2493
2e174807c891 added postprocessed fresh-lemmas for constructors
Christian Urban <urbanc@in.tum.de>
parents: 2492
diff changeset
    35
thm single_let.fresh
2461
86028b2016bd some experiments with support
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    36
2509
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    37
lemmas permute_bn = permute_bn_raw.simps[quot_lifted]
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    38
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    39
lemma uu:
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    40
  shows "alpha_bn as (permute_bn p as)"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    41
apply(induct as rule: single_let.inducts(2))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    42
apply(auto)[7]
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    43
apply(simp add: permute_bn)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    44
apply(simp add: single_let.eq_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    45
done
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    46
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    47
lemma tt:
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    48
  shows "(p \<bullet> bn as) = bn (permute_bn p as)"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    49
apply(induct as rule: single_let.inducts(2))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    50
apply(auto)[7]
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    51
apply(simp add: permute_bn single_let.bn_defs)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    52
apply(simp add: atom_eqvt)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    53
done
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    54
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    55
lemma Abs_fresh_star':
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    56
  fixes x::"'a::fs"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    57
  shows  "set bs = as \<Longrightarrow> as \<sharp>* Abs_lst bs x"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    58
  unfolding fresh_star_def
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    59
  by(simp_all add: Abs_fresh_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    60
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    61
lemma strong_exhaust:
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    62
  fixes c::"'a::fs"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    63
  assumes "\<And>name. y = Var name \<Longrightarrow> P" 
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    64
  and "\<And>trm1 trm2. y = App trm1 trm2 \<Longrightarrow> P" 
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    65
  and "\<And>name trm. \<lbrakk>{atom name} \<sharp>* c; y = Lam name trm\<rbrakk> \<Longrightarrow> P"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    66
  and "\<And>assg trm. \<lbrakk>set (bn assg) \<sharp>* c; y = Let assg trm\<rbrakk> \<Longrightarrow> P"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    67
  and "\<And>name1 name2 trm1 trm2 trm3. \<lbrakk>{atom name1} \<sharp>* c; y = Foo name1 name2 trm1 trm2 trm3\<rbrakk> \<Longrightarrow> P"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    68
  and "\<And>name1 name2 trm. \<lbrakk>{atom name1, atom name2} \<sharp>* c; y = Bar name1 name2 trm\<rbrakk> \<Longrightarrow> P" 
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    69
  and "\<And>name trm1 trm2. \<lbrakk>{atom name} \<sharp>* c; y = Baz name trm1 trm2\<rbrakk> \<Longrightarrow> P"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    70
  shows "P"
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    71
  apply(rule_tac y="y" in single_let.exhaust(1))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    72
  apply(rule assms(1))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    73
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    74
  apply(rule assms(2))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    75
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    76
  apply(subgoal_tac "\<exists>q. (q \<bullet> {atom name}) \<sharp>* c \<and> supp ([[atom name]]lst.trm) \<sharp>* q")
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    77
  apply(erule exE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    78
  apply(erule conjE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    79
  apply(perm_simp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    80
  apply(rule assms(3))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    81
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    82
  apply(drule supp_perm_eq[symmetric])
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    83
  apply(simp add: single_let.eq_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    84
  apply(perm_simp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    85
  apply(rule refl)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    86
  apply(rule at_set_avoiding2)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    87
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    88
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    89
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    90
  apply(simp add: Abs_fresh_star')
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    91
  apply(subgoal_tac "\<exists>q. (q \<bullet> (set (bn assg))) \<sharp>* (c::'a::fs) \<and> supp ([bn assg]lst.trm) \<sharp>* q")
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    92
  apply(erule exE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    93
  apply(erule conjE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    94
  apply(perm_simp add: tt)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    95
  apply(rule_tac assms(4))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    96
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    97
  apply(drule supp_perm_eq[symmetric])
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    98
  apply(simp add: single_let.eq_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    99
  apply(simp add: tt uu)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   100
  apply(rule at_set_avoiding2)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   101
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   102
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   103
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   104
  apply(simp add: Abs_fresh_star)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   105
  apply(subgoal_tac "\<exists>q. (q \<bullet> {atom name1}) \<sharp>* (c::'a::fs) \<and> 
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   106
    supp ([{atom name1}]set.(((name2, trm1), trm2), trm3)) \<sharp>* q")
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   107
  apply(erule exE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   108
  apply(erule conjE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   109
  apply(perm_simp add: tt)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   110
  apply(rule_tac assms(5))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   111
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   112
  apply(drule supp_perm_eq[symmetric])
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   113
  apply(simp add: single_let.eq_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   114
  apply(perm_simp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   115
  apply(rule refl)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   116
  apply(rule at_set_avoiding2)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   117
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   118
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   119
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   120
  apply(simp add: Abs_fresh_star)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   121
  apply(subgoal_tac "\<exists>q. (q \<bullet> {atom name1, atom name2}) \<sharp>* (c::'a::fs) \<and> 
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   122
    supp ([[atom name2, atom name1]]lst.((trm, name1), name2)) \<sharp>* q")
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   123
  apply(erule exE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   124
  apply(erule conjE)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   125
  apply(perm_simp add: tt)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   126
  apply(rule_tac assms(6))
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   127
  apply(assumption)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   128
  apply(drule supp_perm_eq[symmetric])
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   129
  apply(simp add: single_let.eq_iff)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   130
  apply(perm_simp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   131
  apply(rule refl)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   132
  apply(rule at_set_avoiding2)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   133
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   134
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   135
  apply(simp add: finite_supp)
679cef364022 llncs and more sqeezing
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   136
  oops
2524
693562f03eee major reorganisation of fset (renamed fset_to_set to fset, changed the definition of list_eq and fcard_raw)
Christian Urban <urbanc@in.tum.de>
parents: 2509
diff changeset
   137
2064
2725853f43b9 solved the problem with equivariance by first eta-normalising the goal
Christian Urban <urbanc@in.tum.de>
parents: 2028
diff changeset
   138
1596
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
end
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
c69d9fb16785 Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142