Nominal/Abs.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 23 Mar 2010 08:33:48 +0100
changeset 1595 aeed597d2043
parent 1588 7cebb576fae3
child 1657 d7dc35222afc
permissions -rw-r--r--
Move examples which create more permutations out
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
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  supp_abs_fun
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   123
where
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  "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
   125
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   126
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
lemma supp_abs_fun_lemma:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  assumes a: "x \<approx>abs y" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  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
   130
  using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  apply(induct rule: alpha_abs.induct)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   135
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   136
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
   137
  apply(rule equivpI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  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
   139
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  (* refl *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  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
   142
  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
   143
  apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  (* symm *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  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
   147
  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
   148
  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
   149
  apply(clarsimp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  (* trans *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  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
   153
  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
   154
  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
   155
  apply(auto)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   159
  "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
   160
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  "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
   162
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
lemma [quot_respect]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
  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
   165
  apply(clarsimp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
lemma [quot_respect]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  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
   173
  apply(clarsimp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  apply(rule alpha_gen_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  apply(simp_all add: supp_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  done
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 "(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
   181
  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
   182
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
lemma abs_induct:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  "\<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
   186
  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
   187
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
(* TEST case *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
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
   191
thm abs_induct abs_induct2
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   193
instantiation abs_gen :: (pt) pt
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   197
  "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
   198
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  "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
   200
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   201
(* ??? 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
   202
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
   203
  fixes x::"'a::pt"  
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  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
   205
  thm permute_prod.simps
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  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
   207
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
instance
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
  apply(default)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  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
   211
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
quotient_definition
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   217
  "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
   218
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  "supp_abs_fun"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
lemma supp_Abs_fun_simp:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  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
   223
  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
   224
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
lemma supp_Abs_fun_eqvt [eqvt]:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  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
   227
  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
   228
  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
   229
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
lemma supp_Abs_fun_fresh:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  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
   233
  apply(rule fresh_fun_eqvt_app)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  apply(simp add: eqvts_raw)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
lemma Abs_swap:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  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
   240
  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
   241
  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
   242
  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
   243
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
lemma Abs_supports:
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   245
  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
   246
  unfolding supports_def
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  apply(simp (no_asm))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  apply(subst Abs_swap[symmetric])
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   253
lemma finite_supp_Abs_subset1:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   254
  assumes "finite (supp x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  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
   256
  apply(simp add: supp_conv_fresh)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  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
   259
  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
   260
  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
   261
  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
   262
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   264
lemma finite_supp_Abs_subset2:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   265
  assumes "finite (supp x)"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  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
   267
  apply(rule supp_is_subset)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
  apply(rule Abs_supports)
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   269
  apply(simp add: assms)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   270
  done
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   271
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   272
lemma finite_supp_Abs:
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   273
  assumes "finite (supp x)"
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   274
  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
   275
  apply(rule_tac subset_antisym)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   276
  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
   277
  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
   278
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
lemma supp_Abs:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  fixes x::"'a::fs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  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
   283
  apply(rule finite_supp_Abs)
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   284
  apply(simp add: finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   287
instance abs_gen :: (fs) fs
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  apply(default)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  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
   290
  apply(simp add: supp_Abs)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  apply(simp add: finite_supp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
lemma Abs_fresh_iff:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  fixes x::"'a::fs"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
  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
   297
  apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  apply(simp add: supp_Abs)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
lemma Abs_eq_iff:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
  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
   304
  by (lifting alpha_abs.simps(1))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
(* 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
  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
   310
  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
   311
  coincide
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
*)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
fun
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  alpha1
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
where
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
  "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
   318
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
notation 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
  alpha1 ("_ \<approx>abs1 _")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   322
fun
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   323
  alpha2
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   324
where
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   325
  "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
   326
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   327
notation 
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   328
  alpha2 ("_ \<approx>abs2 _")
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   329
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   330
lemma alpha_old_new:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  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
   332
  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
   333
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
apply(erule disjE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
apply(rule exI)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
apply(rule alpha_gen_refl)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
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
   341
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
apply(rule conjI)
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   344
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
   345
apply(rule trans)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
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
   347
apply(subst swap_set_not_in)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
back
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
apply(simp)
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(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
   352
apply(rule conjI)
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   353
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
   354
apply(simp add: permute_self)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
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
   356
apply(simp add: permute_set_eq)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
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
   358
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
   359
apply(blast)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
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
   361
apply(simp add: eqvts)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
lemma perm_induct_test:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
  fixes P :: "perm => bool"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
  assumes fin: "finite (supp p)" 
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
  assumes zero: "P 0"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
  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
   370
  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
   371
  shows "P p"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
using fin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
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
   374
oops
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   375
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   376
lemma ii:
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   377
  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
   378
  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
   379
using assms
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   380
apply(auto)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   381
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
   382
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
   383
done
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   384
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   385
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   386
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   387
lemma alpha_abs_Pair:
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   388
  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
   389
         \<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
   390
  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
   391
  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
   392
  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
   393
  apply(rule iffI)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   394
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   395
  defer
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   396
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   397
  apply(rule conjI)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   398
  apply(clarify)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   399
  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
   400
  apply(rule sym)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   401
  apply(rule ii)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   402
  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
   403
  apply(clarify)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   404
  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
   405
  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
   406
  apply(rule sym)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   407
  apply(rule ii)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   408
  apply(simp)
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   409
  done
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
   410
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
lemma yy:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
  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
   414
  shows "S1 = S2"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
using assms
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
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
   417
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
lemma kk:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
  assumes a: "p \<bullet> x = y"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
  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
   422
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
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
   424
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
   425
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
   426
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
lemma ww:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  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
   430
  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
   431
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
   432
apply(simp add: supports_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
using assms
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
apply -
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
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
   436
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
apply(rule supp_supports)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
apply(auto)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
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
   440
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
   441
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
   442
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
lemma alpha_abs_sym:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  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
   446
  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
   447
using a
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(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
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
   451
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
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
   453
apply (metis permute_minus_cancel(2))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
lemma alpha_abs_trans:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  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
   458
  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
   459
  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
   460
using a b
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
apply(erule exE)+
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
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
   464
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
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
   466
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
lemma alpha_equal:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
  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
   470
  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
   471
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
apply(erule exE)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
apply(erule conjE)+
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
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
   477
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
apply(subgoal_tac "supp x \<sharp>* p")
1563
eb60f360a200 moved lemmas supp_perm_eq and exists_perm to Nominal2_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1558
diff changeset
   479
apply(drule supp_perm_eq)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
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
   484
apply(simp)
1563
eb60f360a200 moved lemmas supp_perm_eq and exists_perm to Nominal2_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1558
diff changeset
   485
apply(drule supp_perm_eq)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
apply(simp (no_asm_use))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
apply(drule yy)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
apply(case_tac "a \<sharp> p")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
apply(subgoal_tac "supp y \<sharp>* p")
1563
eb60f360a200 moved lemmas supp_perm_eq and exists_perm to Nominal2_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1558
diff changeset
   496
apply(drule supp_perm_eq)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
apply(clarify)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
apply(simp (no_asm_use))
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
apply(metis)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
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
   501
apply(frule_tac kk)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
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
   503
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
apply(simp add: supp_perm)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
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
   507
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
   508
apply(simp add: fresh_star_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
lemma alpha_unequal:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
  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
   513
  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
   514
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
apply -
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
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
   517
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
   518
defer
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
apply(simp add: alpha_gen)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
apply(drule_tac alpha_abs_swap)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
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
   524
apply(drule alpha_abs_sym)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
apply(rotate_tac 4)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
apply(drule_tac alpha_abs_trans)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
apply(assumption)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
apply(drule alpha_equal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
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
   530
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
   531
apply(simp add: fresh_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
apply(simp add: fresh_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
lemma alpha_new_old:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
  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
   537
  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
   538
using a
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
apply(case_tac "a=b")
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
apply(simp only: alpha_equal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
apply(drule alpha_unequal)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
apply(simp)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
(* support of concrete atom sets *)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
lemma supp_atom_image:
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  fixes as::"'a::at_base set"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
  shows "supp (atom ` as) = supp as"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
apply(simp add: supp_def)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
apply(simp add: image_eqvt)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
apply(simp add: atom_eqvt_raw)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
apply(simp add: atom_image_cong)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   558
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
   559
  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
   560
  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
   561
  apply (fold fresh_def)
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   562
  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
   563
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
1467
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   565
(* 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
   566
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
   567
  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
   568
  by auto
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   569
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   570
lemma split_prs2[quot_preserve]:
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   571
  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
   572
  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
   573
  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
   574
  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
   575
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   576
lemma alpha_gen2:
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   577
  "(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
   578
  (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
   579
  \<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
   580
by (simp add: alpha_gen)
77b86f1fc936 Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1460
diff changeset
   581
1544
c6849a634582 Keep only one copy of infinite_Un.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1543
diff changeset
   582
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   583
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
   584
  fixes pi
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   585
  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
   586
  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
   587
  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
   588
  using b apply -
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   589
  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
   590
  apply(erule conjE)+
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   591
  apply(rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   592
  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
   593
  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
   594
  apply simp
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   595
  apply(clarify)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   596
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   597
  apply(rule a)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   598
  apply assumption
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   599
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   600
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   601
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
   602
  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
   603
  (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
   604
  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
   605
  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
   606
  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
   607
  using a
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   608
  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
   609
  apply clarify
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   610
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   611
  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
   612
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   613
  apply (rotate_tac 3)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   614
  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
   615
  apply simp
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   616
  apply (rule conjI)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   617
  apply (rotate_tac -1)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   618
  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
   619
  apply simp_all
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   620
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   621
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   622
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
   623
  fixes pi pia
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   624
  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
   625
  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
   626
  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
   627
  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
   628
  using b c apply -
1581
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   629
  apply(simp add: alpha_gen)
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   630
  apply(erule conjE)+
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   631
  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
   632
  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
   633
  apply(drule mp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   634
  apply(rotate_tac 5)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   635
  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
   636
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   637
  apply(rotate_tac 7)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   638
  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
   639
  apply(simp)
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   640
  done
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   641
1581
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   642
lemma alpha_gen_compose_trans2:
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   643
  fixes pi pia
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   644
  assumes b: "(aa, (t1, t2)) \<approx>gen
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   645
    (\<lambda>(b, a) (d, c). R1 b d \<and> (\<forall>z. R1 d z \<longrightarrow> R1 b z) \<and> R2 a c \<and> (\<forall>z. R2 c z \<longrightarrow> R2 a z))
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   646
    (\<lambda>(b, a). fv_a b \<union> fv_b a) pi (ab, (ta1, ta2))"
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   647
  and c: "(ab, (ta1, ta2)) \<approx>gen (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   648
    pia (ac, (sa1, sa2))"
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   649
  and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)"
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   650
  and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)"
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   651
  shows "(aa, (t1, t2)) \<approx>gen (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   652
    (pia + pi) (ac, (sa1, sa2))"
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   653
  using b c apply -
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   654
  apply(simp add: alpha_gen2)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   655
  apply(simp add: alpha_gen)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   656
  apply(erule conjE)+
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   657
  apply(simp add: fresh_star_plus)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   658
  apply(drule_tac x="- pia \<bullet> sa1" in spec)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   659
  apply(drule mp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   660
  apply(rotate_tac 5)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   661
  apply(drule_tac pi="- pia" in r1)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   662
  apply(simp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   663
  apply(rotate_tac -1)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   664
  apply(drule_tac pi="pia" in r1)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   665
  apply(simp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   666
  apply(drule_tac x="- pia \<bullet> sa2" in spec)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   667
  apply(drule mp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   668
  apply(rotate_tac 6)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   669
  apply(drule_tac pi="- pia" in r2)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   670
  apply(simp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   671
  apply(rotate_tac -1)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   672
  apply(drule_tac pi="pia" in r2)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   673
  apply(simp)
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   674
  done
6b1eea8dcdc0 equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1563
diff changeset
   675
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677