Nominal/Abs.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 20 Mar 2010 04:51:26 +0100
changeset 1558 a5ba76208983
parent 1557 fee2389789ad
child 1563 eb60f360a200
permissions -rw-r--r--
started cleaning up and introduced 3 versions of ~~gen
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Abs
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     2
imports "Nominal2_Atoms" 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     3
        "Nominal2_Eqvt" 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     4
        "Nominal2_Supp" 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     5
        "Nominal2_FSet"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     6
        "../Quotient" 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
     7
        "../Quotient_Product" 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  alpha_gen 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
where
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  alpha_gen[simp del]:
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    14
  "alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> 
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    15
     f x - bs = f y - cs \<and> 
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    16
     (f x - bs) \<sharp>* pi \<and> 
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    17
     R (pi \<bullet> x) y \<and>
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    18
     pi \<bullet> bs = cs"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    20
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    21
  alpha_res
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    22
where
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    23
  alpha_res[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    24
  "alpha_res (bs, x) R f pi (cs, y) \<longleftrightarrow> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    25
     f x - bs = f y - cs \<and> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    26
     (f x - bs) \<sharp>* pi \<and> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    27
     R (pi \<bullet> x) y"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    29
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    30
  alpha_lst
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    31
where
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    32
  alpha_lst[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    33
  "alpha_lst (bs, x) R f pi (cs, y) \<longleftrightarrow> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    34
     f x - set bs = f y - set cs \<and> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    35
     (f x - set bs) \<sharp>* pi \<and> 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    36
     R (pi \<bullet> x) y \<and>
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    37
     pi \<bullet> bs = cs"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    38
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    39
lemmas alphas = alpha_gen.simps alpha_res.simps alpha_lst.simps
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    40
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    41
notation
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    42
  alpha_gen ("_ \<approx>gen _ _ _ _" [100, 100, 100, 100, 100] 100) and
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    43
  alpha_res ("_ \<approx>res _ _ _ _" [100, 100, 100, 100, 100] 100) and
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    44
  alpha_lst ("_ \<approx>lst _ _ _ _" [100, 100, 100, 100, 100] 100) 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    45
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    46
(* monos *)
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    47
lemma [mono]: 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    48
  shows "R1 \<le> R2 \<Longrightarrow> alpha_gen bs R1 \<le> alpha_gen bs R2"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    49
  and   "R1 \<le> R2 \<Longrightarrow> alpha_res bs R1 \<le> alpha_res bs R2"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    50
  and   "R1 \<le> R2 \<Longrightarrow> alpha_lst cs R1 \<le> alpha_lst cs R2"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    51
  by (case_tac [!] bs, case_tac [!] cs) 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    52
     (auto simp add: le_fun_def le_bool_def alphas)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
lemma alpha_gen_refl:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  assumes a: "R x x"
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
    56
  shows "(bs, x) \<approx>gen R f 0 (bs, x)"
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    57
  and   "(bs, x) \<approx>res R f 0 (bs, x)"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    58
  and   "(cs, x) \<approx>lst R f 0 (cs, x)"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    59
  using a 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    60
  unfolding alphas
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    61
  unfolding fresh_star_def
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    62
  by (simp_all add: fresh_zero_perm)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
lemma alpha_gen_sym:
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    65
  assumes a: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    66
  shows "(bs, x) \<approx>gen R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>gen R f (- p) (bs, x)"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    67
  and   "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    68
  and   "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    69
  using a
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    70
  unfolding alphas
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    71
  unfolding fresh_star_def
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    72
  by (auto simp add:  fresh_minus_perm)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
lemma alpha_gen_trans:
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    75
  assumes a: "\<lbrakk>R (p \<bullet> x) y; R (q \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((q + p) \<bullet> x) z"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    76
  shows "\<lbrakk>(bs, x) \<approx>gen R f p (cs, y); (cs, y) \<approx>gen R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>gen R f (q + p) (ds, z)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    77
  and   "\<lbrakk>(bs, x) \<approx>res R f p (cs, y); (cs, y) \<approx>res R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>res R f (q + p) (ds, z)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    78
  and   "\<lbrakk>(es, x) \<approx>lst R f p (gs, y); (gs, y) \<approx>lst R f q (hs, z)\<rbrakk> \<Longrightarrow> (es, x) \<approx>lst R f (q + p) (hs, z)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    79
  using a 
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    80
  unfolding alphas
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    81
  unfolding fresh_star_def
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    82
  by (simp_all add: fresh_plus_perm)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
lemma alpha_gen_eqvt:
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    85
  assumes a: "R (q \<bullet> x) y \<Longrightarrow> R (p \<bullet> (q \<bullet> x)) (p \<bullet> y)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    86
  and     b: "p \<bullet> (f x) = f (p \<bullet> x)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    87
  and     c: "p \<bullet> (f y) = f (p \<bullet> y)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    88
  shows "(bs, x) \<approx>gen R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>gen R f (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    89
  and   "(bs, x) \<approx>res R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>res R f (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    90
  and   "(ds, x) \<approx>lst R f q (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>lst R f (p \<bullet> q) (p \<bullet> es, p \<bullet> y)" 
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    91
  unfolding alphas
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    92
  unfolding set_eqvt[symmetric]
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    93
  unfolding b[symmetric] c[symmetric]
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    94
  unfolding Diff_eqvt[symmetric]
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    95
  unfolding permute_eqvt[symmetric]
1487
b55b78e63913 Proper compose_sym2
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1482
diff changeset
    96
  using a
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
    97
  by (auto simp add: fresh_star_permute_iff)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  alpha_abs 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
where
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   102
  "alpha_abs (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
notation
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  alpha_abs ("_ \<approx>abs _")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
lemma alpha_abs_swap:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  assumes a1: "a \<notin> (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  and     a2: "b \<notin> (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  shows "(bs, x) \<approx>abs ((a \<rightleftharpoons> b) \<bullet> bs, (a \<rightleftharpoons> b) \<bullet> x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  using a1 a2
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   112
  unfolding Diff_iff
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   113
  unfolding alpha_abs.simps
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   114
  unfolding alphas
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   115
  unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric]
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   116
  unfolding fresh_star_def fresh_def
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   117
  unfolding swap_set_not_in[OF a1 a2] 
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   118
  by (rule_tac x="(a \<rightleftharpoons> b)" in exI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   119
     (auto simp add: supp_perm swap_atom)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   121
lemma alpha_gen_swap_fun:
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   122
  assumes f_eqvt: "\<And>pi. (pi \<bullet> (f x)) = f (pi \<bullet> x)"
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   123
  assumes a1: "a \<notin> (f x) - bs"
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   124
  and     a2: "b \<notin> (f x) - bs"
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   125
  shows "\<exists>pi. (bs, x) \<approx>gen (op=) f pi ((a \<rightleftharpoons> b) \<bullet> bs, (a \<rightleftharpoons> b) \<bullet> x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  apply(rule_tac x="(a \<rightleftharpoons> b)" in exI)
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   127
  apply(simp add: alpha_gen)
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   128
  apply(simp add: f_eqvt[symmetric] Diff_eqvt[symmetric])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  apply(simp add: swap_set_not_in[OF a1 a2])
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  using a1 a2
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  apply(simp add: fresh_star_def fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  apply(blast)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  apply(simp add: supp_swap)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  supp_abs_fun
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   139
where
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  "supp_abs_fun (bs, x) = (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   142
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
lemma supp_abs_fun_lemma:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  assumes a: "x \<approx>abs y" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  shows "supp_abs_fun x = supp_abs_fun y"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  apply(induct rule: alpha_abs.induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   151
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   152
quotient_type 'a abs_gen = "(atom set \<times> 'a::pt)" / "alpha_abs"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  apply(rule equivpI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  unfolding reflp_def symp_def transp_def
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  (* refl *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  apply(clarify)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   158
  apply(rule_tac x="0" in exI)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  (* symm *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  apply(clarify)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   163
  apply(rule_tac x="- p" in exI)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
  apply(rule alpha_gen_sym)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   165
  apply(clarsimp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  (* trans *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  apply(clarify)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   169
  apply(rule_tac x="pa + p" in exI)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  apply(rule alpha_gen_trans)
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   171
  apply(auto)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   175
  "Abs::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_gen"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
lemma [quot_respect]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
  shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  apply(clarsimp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
lemma [quot_respect]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  apply(clarsimp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  apply(rule alpha_gen_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  apply(simp_all add: supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
lemma [quot_respect]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  apply(simp add: supp_abs_fun_lemma)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
lemma abs_induct:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  "\<lbrakk>\<And>as (x::'a::pt). P (Abs as x)\<rbrakk> \<Longrightarrow> P t"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  apply(lifting prod.induct[where 'a="atom set" and 'b="'a"])
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
(* TEST case *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
lemmas abs_induct2 = prod.induct[where 'a="atom set" and 'b="'a::pt", quot_lifted]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
thm abs_induct abs_induct2
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   209
instantiation abs_gen :: (pt) pt
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   213
  "permute_abs_gen::perm \<Rightarrow> ('a::pt abs_gen) \<Rightarrow> 'a abs_gen"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   217
(* ??? has to be 'a \<dots> 'b does not work *)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
lemma permute_ABS [simp]:
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   219
  fixes x::"'a::pt"  
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  shows "(p \<bullet> (Abs as x)) = Abs (p \<bullet> as) (p \<bullet> x)"
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   221
  thm permute_prod.simps
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"])
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
instance
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
  apply(default)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  apply(induct_tac [!] x rule: abs_induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   233
  "supp_Abs_fun :: ('a::pt) abs_gen \<Rightarrow> atom \<Rightarrow> bool"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  "supp_abs_fun"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
lemma supp_Abs_fun_simp:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  shows "supp_Abs_fun (Abs bs x) = (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  by (lifting supp_abs_fun.simps(1))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
lemma supp_Abs_fun_eqvt [eqvt]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  shows "(p \<bullet> supp_Abs_fun x) = supp_Abs_fun (p \<bullet> x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  apply(induct_tac x rule: abs_induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
lemma supp_Abs_fun_fresh:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_Abs_fun (Abs bs x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  apply(rule fresh_fun_eqvt_app)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  apply(simp add: eqvts_raw)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
lemma Abs_swap:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  assumes a1: "a \<notin> (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  and     a2: "b \<notin> (supp x) - bs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  shows "(Abs bs x) = (Abs ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x))"
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   258
  using a1 a2 by (lifting alpha_abs_swap)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
lemma Abs_supports:
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   261
  shows "((supp x) - as) supports (Abs as x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  unfolding supports_def
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  apply(simp (no_asm))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  apply(subst Abs_swap[symmetric])
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   269
lemma finite_supp_Abs_subset1:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   270
  assumes "finite (supp x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  shows "(supp x) - as \<subseteq> supp (Abs as x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  apply(simp add: supp_conv_fresh)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  apply(drule_tac supp_Abs_fun_fresh)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  apply(simp only: supp_Abs_fun_simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  apply(simp add: fresh_def)
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   277
  apply(simp add: supp_finite_atom_set assms)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   280
lemma finite_supp_Abs_subset2:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   281
  assumes "finite (supp x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  shows "supp (Abs as x) \<subseteq> (supp x) - as"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  apply(rule supp_is_subset)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  apply(rule Abs_supports)
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   285
  apply(simp add: assms)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   286
  done
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   287
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   288
lemma finite_supp_Abs:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   289
  assumes "finite (supp x)"
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   290
  shows "supp (Abs as x) = (supp x) - as"
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   291
  apply(rule_tac subset_antisym)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   292
  apply(rule finite_supp_Abs_subset2[OF assms])
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   293
  apply(rule finite_supp_Abs_subset1[OF assms])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
lemma supp_Abs:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
  fixes x::"'a::fs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  shows "supp (Abs as x) = (supp x) - as"
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   299
  apply(rule finite_supp_Abs)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   300
  apply(simp add: finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   303
instance abs_gen :: (fs) fs
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  apply(default)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
  apply(induct_tac x rule: abs_induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  apply(simp add: supp_Abs)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  apply(simp add: finite_supp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
lemma Abs_fresh_iff:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
  fixes x::"'a::fs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
  shows "a \<sharp> Abs bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  apply(simp add: supp_Abs)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
lemma Abs_eq_iff:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
  shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
  by (lifting alpha_abs.simps(1))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
(* 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  below is a construction site for showing that in the
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
  single-binder case, the old and new alpha equivalence 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
  coincide
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
*)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  alpha1
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
where
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  "alpha1 (a, x) (b, y) \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<rightleftharpoons> b) \<bullet> y \<and> a \<sharp> y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
notation 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  alpha1 ("_ \<approx>abs1 _")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   338
fun
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   339
  alpha2
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   340
where
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   341
  "alpha2 (a, x) (b, y) \<longleftrightarrow> (\<exists>c. c \<sharp> (a,b,x,y) \<and> ((c \<rightleftharpoons> a) \<bullet> x) = ((c \<rightleftharpoons> b) \<bullet> y))"
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   342
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   343
notation 
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   344
  alpha2 ("_ \<approx>abs2 _")
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   345
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   346
lemma alpha_old_new:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  assumes a: "(a, x) \<approx>abs1 (b, y)" "sort_of a = sort_of b"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
  shows "({a}, x) \<approx>abs ({b}, y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
apply(erule disjE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
apply(rule_tac x="(a \<rightleftharpoons> b)" in  exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
apply(rule conjI)
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   360
apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in  permute_eq_iff[THEN iffD1])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
apply(rule trans)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
apply(simp add: Diff_eqvt supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
apply(subst swap_set_not_in)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
back
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
apply(simp add: permute_set_eq)
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   368
apply(rule conjI)
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   369
apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in fresh_star_permute_iff[THEN iffD1])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
apply(simp add: permute_self)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
apply(simp add: Diff_eqvt supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
apply(simp add: permute_set_eq)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
apply(simp add: fresh_star_def fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
apply(blast)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
apply(simp add: supp_swap)
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   377
apply(simp add: eqvts)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
lemma perm_zero:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  assumes a: "\<forall>x::atom. p \<bullet> x = x"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
  shows "p = 0"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
by (simp add: expand_perm_eq)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
  add_perm 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
where
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
  "add_perm [] = 0"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
| "add_perm ((a, b) # xs) = (a \<rightleftharpoons> b) + add_perm xs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
  elem_perm
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
where
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
  "elem_perm [] = {}"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
| "elem_perm ((a, b) # xs) = {a, b} \<union> elem_perm xs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
lemma add_perm_apend:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
  shows "add_perm (xs @ ys) = add_perm xs + add_perm ys"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
apply(induct xs arbitrary: ys)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
apply(auto simp add: add_assoc)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
lemma perm_list_exists:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
  fixes p::perm
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
  shows "\<exists>xs. p = add_perm xs \<and> supp xs \<subseteq> supp p"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
apply(induct p taking: "\<lambda>p::perm. card (supp p)" rule: measure_induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
apply(rename_tac p)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
apply(case_tac "supp p = {}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
apply(simp add: supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
apply(drule perm_zero)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
apply(rule_tac x="[]" in exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
apply(simp add: supp_Nil)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
apply(subgoal_tac "\<exists>x. x \<in> supp p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
apply(auto)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
apply(drule_tac x="p + (((- p) \<bullet> x) \<rightleftharpoons> x)" in spec)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
apply(drule mp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
apply(rule_tac x="xs @ [((- p) \<bullet> x, x)]" in exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
apply(simp add: add_perm_apend)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
apply(erule conjE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
apply(drule sym)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
apply(simp add: expand_perm_eq)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
apply(simp add: supp_append)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
apply(simp add: supp_perm supp_Cons supp_Nil supp_Pair supp_atom)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
apply(rule conjI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
prefer 2
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
apply(auto)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
apply (metis left_minus minus_unique permute_atom_def_raw permute_minus_cancel(2))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
apply(rule psubset_card_mono)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
apply(simp add: finite_supp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
apply(rule psubsetI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
apply(subgoal_tac "x \<notin> supp (p + (- p \<bullet> x \<rightleftharpoons> x))")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
apply(blast)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
apply(simp add: supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
apply(auto simp add: supp_perm)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
apply(case_tac "x = xa")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
apply(case_tac "((- p) \<bullet> x) = xa")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
apply(case_tac "sort_of xa = sort_of x")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
apply(auto)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
apply(subgoal_tac "{a. p \<bullet> (- p \<bullet> x \<rightleftharpoons> x) \<bullet> a \<noteq> a} \<subseteq> {a. p \<bullet> a \<noteq> a}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
apply(blast)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
apply(auto simp add: supp_perm)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
apply(case_tac "x = xa")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
apply(case_tac "((- p) \<bullet> x) = xa")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
apply(case_tac "sort_of xa = sort_of x")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
apply(auto)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
lemma tt0:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
  fixes p::perm
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
  shows "(supp x) \<sharp>* p \<Longrightarrow> \<forall>a \<in> supp p. a \<sharp> x"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
apply(auto simp add: fresh_star_def supp_perm fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
lemma uu0:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  shows "(\<forall>a \<in> elem_perm xs. a \<sharp> x) \<Longrightarrow> (add_perm xs \<bullet> x) = x"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
apply(induct xs rule: add_perm.induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
apply(simp add: swap_fresh_fresh)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
lemma yy0:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
  fixes xs::"(atom \<times> atom) list"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
  shows "supp xs = elem_perm xs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
apply(induct xs rule: elem_perm.induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
apply(auto simp add: supp_Nil supp_Cons supp_Pair supp_atom)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
lemma tt1:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  shows "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
apply(drule tt0)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
apply(subgoal_tac "\<exists>xs. p = add_perm xs \<and> supp xs \<subseteq> supp p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
prefer 2
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
apply(rule perm_list_exists)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
apply(simp only: yy0)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
apply(rule uu0)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
lemma perm_induct_test:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
  fixes P :: "perm => bool"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
  assumes fin: "finite (supp p)" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
  assumes zero: "P 0"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
  assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
  assumes plus: "\<And>p1 p2. \<lbrakk>supp p1 \<inter> supp p2 = {}; P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
  shows "P p"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
using fin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
apply(induct F\<equiv>"supp p" arbitrary: p rule: finite_induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
apply(simp add: supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
apply(drule perm_zero)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
apply(simp add: zero)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
apply(rotate_tac 3)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
oops
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   517
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   518
lemma ii:
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   519
  assumes "\<forall>x \<in> A. p \<bullet> x = x"
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   520
  shows "p \<bullet> A = A"
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   521
using assms
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   522
apply(auto)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   523
apply (metis Collect_def Collect_mem_eq Int_absorb assms eqvt_bound inf_Int_eq mem_def mem_permute_iff)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   524
apply (metis Collect_def Collect_mem_eq Int_absorb assms eqvt_apply eqvt_bound eqvt_lambda inf_Int_eq mem_def mem_permute_iff permute_minus_cancel(2) permute_pure)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   525
done
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   526
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   527
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   528
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   529
lemma alpha_abs_Pair:
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   530
  shows "(bs, (x1, x2)) \<approx>gen (\<lambda>(x1,x2) (y1,y2). x1=y1 \<and> x2=y2) (\<lambda>(x,y). supp x \<union> supp y) p (cs, (y1, y2))
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   531
         \<longleftrightarrow> ((bs, x1) \<approx>gen (op=) supp p (cs, y1) \<and> (bs, x2) \<approx>gen (op=) supp p (cs, y2))"         
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   532
  apply(simp add: alpha_gen)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   533
  apply(simp add: fresh_star_def)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   534
  apply(simp add: ball_Un Un_Diff)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   535
  apply(rule iffI)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   536
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   537
  defer
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   538
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   539
  apply(rule conjI)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   540
  apply(clarify)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   541
  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   542
  apply(rule sym)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   543
  apply(rule ii)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   544
  apply(simp add: fresh_def supp_perm)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   545
  apply(clarify)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   546
  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   547
  apply(simp add: fresh_def supp_perm)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   548
  apply(rule sym)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   549
  apply(rule ii)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   550
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   551
  done
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   552
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
lemma yy:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  assumes "S1 - {x} = S2 - {x}" "x \<in> S1" "x \<in> S2"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  shows "S1 = S2"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
using assms
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
apply (metis insert_Diff_single insert_absorb)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
lemma kk:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
  assumes a: "p \<bullet> x = y"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
  shows "\<forall>a \<in> supp x. (p \<bullet> a) \<in> supp y"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
apply(auto)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   566
apply(rule_tac p="- p" in permute_boolE)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
apply(simp add: mem_eqvt supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
lemma ww:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
  assumes "a \<notin> supp x" "b \<in> supp x" "a \<noteq> b" "sort_of a = sort_of b"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
  shows "((a \<rightleftharpoons> b) \<bullet> x) \<noteq> x"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
apply(subgoal_tac "(supp x) supports x")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
apply(simp add: supports_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
using assms
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
apply -
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
apply(drule_tac x="a" in spec)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
apply(rule supp_supports)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
apply(rotate_tac 1)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   582
apply(drule_tac p="(a \<rightleftharpoons> b)" in permute_boolI)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
apply(simp add: mem_eqvt supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
lemma alpha_abs_sym:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
  assumes a: "({a}, x) \<approx>abs ({b}, y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  shows "({b}, y) \<approx>abs ({a}, x)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
apply(rule_tac x="- p" in exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
apply(simp add: fresh_star_def fresh_minus_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
apply (metis permute_minus_cancel(2))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
lemma alpha_abs_trans:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  assumes a: "({a1}, x1) \<approx>abs ({a2}, x2)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
  assumes b: "({a2}, x2) \<approx>abs ({a3}, x3)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
  shows "({a1}, x1) \<approx>abs ({a3}, x3)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
using a b
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
apply(erule exE)+
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
apply(rule_tac x="pa + p" in exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
apply(simp add: fresh_star_def fresh_plus_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
lemma alpha_equal:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
  assumes a: "({a}, x) \<approx>abs ({a}, y)" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
  shows "(a, x) \<approx>abs1 (a, y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
apply(erule conjE)+
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
apply(case_tac "a \<notin> supp x")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
apply(subgoal_tac "supp x \<sharp>* p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
apply(drule tt1)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
apply(case_tac "a \<notin> supp y")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
apply(drule tt1)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
apply(simp (no_asm_use))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
apply(drule yy)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
apply(case_tac "a \<sharp> p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
apply(subgoal_tac "supp y \<sharp>* p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
apply(drule tt1)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
apply(simp (no_asm_use))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
apply(metis)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
apply(auto simp add: fresh_star_def)[1]
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
apply(frule_tac kk)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
apply(drule_tac x="a" in bspec)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   646
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
apply(simp add: supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
apply(subgoal_tac "((p \<bullet> a) \<sharp> p)")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
apply(simp add: fresh_def supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
apply(simp add: fresh_star_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
lemma alpha_unequal:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b" "a \<noteq> b"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
  shows "(a, x) \<approx>abs1 (b, y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
apply -
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
apply(subgoal_tac "a \<notin> supp x - {a}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
apply(subgoal_tac "b \<notin> supp x - {a}")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
apply(drule_tac alpha_abs_swap)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
apply(simp only: insert_eqvt empty_eqvt swap_atom_simps)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
apply(drule alpha_abs_sym)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
apply(rotate_tac 4)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
apply(drule_tac alpha_abs_trans)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
apply(drule alpha_equal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
apply(simp)
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
   672
apply(rule_tac p="(a \<rightleftharpoons> b)" in permute_boolE)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
apply(simp add: fresh_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
lemma alpha_new_old:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
  shows "(a, x) \<approx>abs1 (b, y)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
apply(case_tac "a=b")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
apply(simp only: alpha_equal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
apply(drule alpha_unequal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
(* support of concrete atom sets *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
lemma supp_atom_image:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
  fixes as::"'a::at_base set"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
  shows "supp (atom ` as) = supp as"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
apply(simp add: supp_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
apply(simp add: image_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
apply(simp add: atom_eqvt_raw)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
apply(simp add: atom_image_cong)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   700
lemma swap_atom_image_fresh: "\<lbrakk>a \<sharp> atom ` (fn :: ('a :: at_base set)); b \<sharp> atom ` fn\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> fn = fn"
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   701
  apply (simp add: fresh_def)
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   702
  apply (simp add: supp_atom_image)
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   703
  apply (fold fresh_def)
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   704
  apply (simp add: swap_fresh_fresh)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
1467
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   707
(* TODO: The following lemmas can be moved somewhere... *)
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   708
lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===>
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   709
  prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split"
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   710
  by auto
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   711
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   712
lemma split_prs2[quot_preserve]:
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   713
  assumes q1: "Quotient R1 Abs1 Rep1"
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   714
  and q2: "Quotient R2 Abs2 Rep2"
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   715
  shows "((Abs1 ---> Abs2 ---> prod_fun Abs1 Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> prod_fun Rep1 Rep2 ---> id) split = split"
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   716
  by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   717
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   718
lemma alpha_gen2:
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   719
  "(bs, x1, x2) \<approx>gen (\<lambda>(x1, y1) (x2, y2). R1 x1 x2 \<and> R2 y1 y2) (\<lambda>(a, b). f1 a \<union> f2 b) pi (cs, y1, y2) =
1470
3127c75275a6 Fix for the change of alpha_gen.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1469
diff changeset
   720
  (f1 x1 \<union> f2 x2 - bs = f1 y1 \<union> f2 y2 - cs \<and> (f1 x1 \<union> f2 x2 - bs) \<sharp>* pi \<and> R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2
3127c75275a6 Fix for the change of alpha_gen.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1469
diff changeset
   721
  \<and> pi \<bullet> bs = cs)"
1467
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   722
by (simp add: alpha_gen)
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   723
1544
c6849a634582 Keep only one copy of infinite_Un.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1543
diff changeset
   724
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   725
lemma alpha_gen_compose_sym:
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   726
  fixes pi
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   727
  assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   728
  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   729
  shows "(ab, s) \<approx>gen R f (- pi) (aa, t)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   730
  using b apply -
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   731
  apply(simp add: alpha_gen)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   732
  apply(erule conjE)+
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   733
  apply(rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   734
  apply(simp add: fresh_star_def fresh_minus_perm)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   735
  apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))")
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   736
  apply simp
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   737
  apply(clarify)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   738
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   739
  apply(rule a)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   740
  apply assumption
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   741
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   742
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   743
lemma alpha_gen_compose_sym2:
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   744
  assumes a: "(aa, t1, t2) \<approx>gen (\<lambda>(x11, x12) (x21, x22).
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   745
  (R1 x11 x21 \<and> R1 x21 x11) \<and> R2 x12 x22 \<and> R2 x22 x12) (\<lambda>(b, a). fb b \<union> fa a) pi (ab, s1, s2)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   746
  and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   747
  and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   748
  shows "(ab, s1, s2) \<approx>gen (\<lambda>(a, b) (d, c). R1 a d \<and> R2 b c) (\<lambda>(b, a). fb b \<union> fa a) (- pi) (aa, t1, t2)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   749
  using a
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   750
  apply(simp add: alpha_gen)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   751
  apply clarify
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   752
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   753
  apply(simp add: fresh_star_def fresh_minus_perm)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   754
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   755
  apply (rotate_tac 3)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   756
  apply (drule_tac pi="- pi" in r1)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   757
  apply simp
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   758
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   759
  apply (rotate_tac -1)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   760
  apply (drule_tac pi="- pi" in r2)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   761
  apply simp_all
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   762
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   763
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   764
lemma alpha_gen_compose_trans:
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   765
  fixes pi pia
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   766
  assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   767
  and c: "(ab, ta) \<approx>gen R f pia (ac, sa)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   768
  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   769
  shows "(aa, t) \<approx>gen R f (pia + pi) (ac, sa)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   770
  using b c apply -
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   771
  apply(simp add: alpha_gen.simps)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   772
  apply(erule conjE)+
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   773
  apply(simp add: fresh_star_plus)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   774
  apply(drule_tac x="- pia \<bullet> sa" in spec)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   775
  apply(drule mp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   776
  apply(rotate_tac 5)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   777
  apply(drule_tac pi="- pia" in a)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   778
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   779
  apply(rotate_tac 7)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   780
  apply(drule_tac pi="pia" in a)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   781
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   782
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   783
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   784
lemma alpha_gen_compose_eqvt:
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   785
  fixes  pia
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   786
  assumes b: "(g d, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2)) f pia (g e, s)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   787
  and     c: "\<And>y. pi \<bullet> (g y) = g (pi \<bullet> y)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   788
  and     a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   789
  shows  "(g (pi \<bullet> d), pi \<bullet> t) \<approx>gen R f (pi \<bullet> pia) (g (pi \<bullet> e), pi \<bullet> s)"
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   790
  using b
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   791
  apply -
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   792
  apply(simp add: alpha_gen.simps)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   793
  apply(erule conjE)+
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   794
  apply(rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   795
  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   796
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric])
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   797
  apply(rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   798
  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   799
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric])
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   800
  apply(subst permute_eqvt[symmetric])
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   801
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   802
  sorry
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   803
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805