Nominal/Nominal2_Abs.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 12 Jul 2012 10:11:32 +0100
changeset 3191 0440bc1a2438
parent 3183 313e6f2cdd89
child 3192 14c7d7e29c44
permissions -rw-r--r--
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2568
8193bbaa07fe merged Nominal-General directory into Nominal; renamed Abs.thy to Nominal2_Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
     1
theory Nominal2_Abs
8193bbaa07fe merged Nominal-General directory into Nominal; renamed Abs.thy to Nominal2_Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
     2
imports "Nominal2_Base" 
2635
64b4cb2c2bf8 simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents: 2616
diff changeset
     3
        "~~/src/HOL/Quotient" 
64b4cb2c2bf8 simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents: 2616
diff changeset
     4
        "~~/src/HOL/Library/Quotient_List"
64b4cb2c2bf8 simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents: 2616
diff changeset
     5
        "~~/src/HOL/Library/Quotient_Product" 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
2473
a3711f07449b added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents: 2471
diff changeset
     8
a3711f07449b added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents: 2471
diff changeset
     9
section {* Abstractions *}
a3711f07449b added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents: 2471
diff changeset
    10
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
fun
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    12
  alpha_set 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
where
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    14
  alpha_set[simp del]:
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    15
  "alpha_set (bs, x) R f pi (cs, y) \<longleftrightarrow> 
1465
4de35639fef0 added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1460
diff changeset
    16
     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
    17
     (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
    18
     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
    19
     pi \<bullet> bs = cs"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    21
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    22
  alpha_res
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    23
where
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[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    25
  "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
    26
     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
    27
     (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
    28
     R (pi \<bullet> x) y"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    30
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    31
  alpha_lst
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    32
where
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[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    34
  "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
    35
     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
    36
     (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
    37
     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
    38
     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
    39
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    40
lemmas alphas = alpha_set.simps alpha_res.simps alpha_lst.simps
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    41
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    42
notation
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    43
  alpha_set ("_ \<approx>set _ _ _ _" [100, 100, 100, 100, 100] 100) and
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    44
  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
    45
  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
    46
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    47
section {* Mono *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    48
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    49
lemma [mono]: 
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    50
  shows "R1 \<le> R2 \<Longrightarrow> alpha_set bs R1 \<le> alpha_set bs R2"
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    51
  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
    52
  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
    53
  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
    54
     (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
    55
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    56
section {* Equivariance *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    57
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    58
lemma alpha_eqvt[eqvt]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    59
  shows "(bs, x) \<approx>set R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>set (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    60
  and   "(bs, x) \<approx>res R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>res (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    61
  and   "(ds, x) \<approx>lst R f q (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>lst (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> es, p \<bullet> y)" 
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    62
  unfolding alphas
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    63
  unfolding permute_eqvt[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    64
  unfolding set_eqvt[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    65
  unfolding permute_fun_app_eq[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    66
  unfolding Diff_eqvt[symmetric]
3004
c6af56de923d more on paper
Christian Urban <urbanc@in.tum.de>
parents: 2943
diff changeset
    67
  unfolding eq_eqvt[symmetric]
c6af56de923d more on paper
Christian Urban <urbanc@in.tum.de>
parents: 2943
diff changeset
    68
  unfolding fresh_star_eqvt[symmetric]
c6af56de923d more on paper
Christian Urban <urbanc@in.tum.de>
parents: 2943
diff changeset
    69
  by (auto simp add: permute_bool_def)
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    70
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    71
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    72
section {* Equivalence *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    73
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    74
lemma alpha_refl:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    75
  assumes a: "R x x"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    76
  shows "(bs, x) \<approx>set R f 0 (bs, x)"
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    77
  and   "(bs, x) \<approx>res R f 0 (bs, x)"
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    78
  and   "(cs, x) \<approx>lst R f 0 (cs, x)"
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    79
  using a 
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    80
  unfolding alphas
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    81
  unfolding fresh_star_def
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    82
  by (simp_all add: fresh_zero_perm)
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    83
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    84
lemma alpha_sym:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    85
  assumes a: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    86
  shows "(bs, x) \<approx>set R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>set R f (- p) (bs, x)"
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    87
  and   "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)"
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    88
  and   "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)"
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    89
  unfolding alphas fresh_star_def
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    90
  using a
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    91
  by (auto simp add:  fresh_minus_perm)
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    92
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    93
lemma alpha_trans:
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    94
  assumes a: "\<lbrakk>R (p \<bullet> x) y; R (q \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((q + p) \<bullet> x) z"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    95
  shows "\<lbrakk>(bs, x) \<approx>set R f p (cs, y); (cs, y) \<approx>set R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>set R f (q + p) (ds, z)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    96
  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)"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    97
  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)"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    98
  using a
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    99
  unfolding alphas fresh_star_def
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   100
  by (simp_all add: fresh_plus_perm)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   101
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   102
lemma alpha_sym_eqvt:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   103
  assumes a: "R (p \<bullet> x) y \<Longrightarrow> R y (p \<bullet> x)" 
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   104
  and     b: "p \<bullet> R = R"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   105
  shows "(bs, x) \<approx>set R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>set R f (- p) (bs, x)" 
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   106
  and   "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)" 
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   107
  and   "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   108
apply(auto intro!: alpha_sym)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   109
apply(drule_tac [!] a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   110
apply(rule_tac [!] p="p" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   111
apply(perm_simp add: permute_minus_cancel b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   112
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   113
apply(perm_simp add: permute_minus_cancel b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   114
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   115
apply(perm_simp add: permute_minus_cancel b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   116
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   117
done
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   118
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   119
lemma alpha_set_trans_eqvt:
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   120
  assumes b: "(cs, y) \<approx>set R f q (ds, z)"
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   121
  and     a: "(bs, x) \<approx>set R f p (cs, y)" 
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   122
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   123
  and     c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   124
  shows "(bs, x) \<approx>set R f (q + p) (ds, z)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   125
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   126
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   127
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   128
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   129
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   130
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   131
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   132
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   133
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   134
apply(drule_tac p="q" in permute_boolI)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   135
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   136
apply(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   137
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   138
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   139
lemma alpha_res_trans_eqvt:
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   140
  assumes  b: "(cs, y) \<approx>res R f q (ds, z)"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   141
  and     a: "(bs, x) \<approx>res R f p (cs, y)" 
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   142
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   143
  and     c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   144
  shows "(bs, x) \<approx>res R f (q + p) (ds, z)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   145
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   146
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   147
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   148
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   149
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   150
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   151
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   152
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   153
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   154
apply(drule_tac p="q" in permute_boolI)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   155
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   156
apply(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   157
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   158
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   159
lemma alpha_lst_trans_eqvt:
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   160
  assumes b: "(cs, y) \<approx>lst R f q (ds, z)"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   161
  and     a: "(bs, x) \<approx>lst R f p (cs, y)" 
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   162
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   163
  and     c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   164
  shows "(bs, x) \<approx>lst R f (q + p) (ds, z)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   165
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   166
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   167
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   168
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   169
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   170
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   171
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   172
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   173
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   174
apply(drule_tac p="q" in permute_boolI)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   175
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   176
apply(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   177
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   178
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   179
lemmas alpha_trans_eqvt = alpha_set_trans_eqvt alpha_res_trans_eqvt alpha_lst_trans_eqvt
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   180
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   181
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   182
section {* General Abstractions *}
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   183
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
fun
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   185
  alpha_abs_set 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   187
  [simp del]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   188
  "alpha_abs_set (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op=) supp p (cs, y))"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   190
fun
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   191
  alpha_abs_lst
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   192
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   193
  [simp del]:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   194
  "alpha_abs_lst (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>lst (op=) supp p (cs, y))"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   195
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   196
fun
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   197
  alpha_abs_res
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   198
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   199
  [simp del]:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   200
  "alpha_abs_res (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op=) supp p (cs, y))"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   201
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
notation
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   203
  alpha_abs_set (infix "\<approx>abs'_set" 50) and
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   204
  alpha_abs_lst (infix "\<approx>abs'_lst" 50) and
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   205
  alpha_abs_res (infix "\<approx>abs'_res" 50)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   206
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   207
lemmas alphas_abs = alpha_abs_set.simps alpha_abs_res.simps alpha_abs_lst.simps
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   208
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   209
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   210
lemma alphas_abs_refl:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   211
  shows "(bs, x) \<approx>abs_set (bs, x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   212
  and   "(bs, x) \<approx>abs_res (bs, x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   213
  and   "(cs, x) \<approx>abs_lst (cs, x)" 
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   214
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   215
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   216
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   217
  by (rule_tac [!] x="0" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   218
     (simp_all add: fresh_zero_perm)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   219
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   220
lemma alphas_abs_sym:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   221
  shows "(bs, x) \<approx>abs_set (cs, y) \<Longrightarrow> (cs, y) \<approx>abs_set (bs, x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   222
  and   "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (cs, y) \<approx>abs_res (bs, x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   223
  and   "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (es, y) \<approx>abs_lst (ds, x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   224
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   225
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   226
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   227
  by (erule_tac [!] exE, rule_tac [!] x="-p" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   228
     (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
   229
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   230
lemma alphas_abs_trans:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   231
  shows "\<lbrakk>(bs, x) \<approx>abs_set (cs, y); (cs, y) \<approx>abs_set (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>abs_set (ds, z)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   232
  and   "\<lbrakk>(bs, x) \<approx>abs_res (cs, y); (cs, y) \<approx>abs_res (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>abs_res (ds, z)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   233
  and   "\<lbrakk>(es, x) \<approx>abs_lst (gs, y); (gs, y) \<approx>abs_lst (hs, z)\<rbrakk> \<Longrightarrow> (es, x) \<approx>abs_lst (hs, z)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   234
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   235
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   236
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   237
  apply(erule_tac [!] exE, erule_tac [!] exE)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   238
  apply(rule_tac [!] x="pa + p" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   239
  by (simp_all add: fresh_plus_perm)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   240
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   241
lemma alphas_abs_eqvt:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   242
  shows "(bs, x) \<approx>abs_set (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs_set (p \<bullet> cs, p \<bullet> y)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   243
  and   "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs_res (p \<bullet> cs, p \<bullet> y)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   244
  and   "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>abs_lst (p \<bullet> es, p \<bullet> y)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   245
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   246
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   247
  unfolding set_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   248
  unfolding supp_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   249
  unfolding Diff_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   250
  apply(erule_tac [!] exE)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   251
  apply(rule_tac [!] x="p \<bullet> pa" in exI)
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3172
diff changeset
   252
  by (auto simp only: fresh_star_permute_iff permute_eqvt[symmetric])
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   253
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   254
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   255
section {* Strengthening the equivalence *}
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   256
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   257
lemma disjoint_right_eq:
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   258
  assumes a: "A \<union> B1 = A \<union> B2"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   259
  and     b: "A \<inter> B1 = {}" "A \<inter> B2 = {}"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   260
  shows "B1 = B2"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   261
using a b
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   262
by (metis Int_Un_distrib2 Int_absorb2 Int_commute Un_upper2)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   263
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   264
lemma supp_property_res:
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   265
  assumes a: "(as, x) \<approx>res (op =) supp p (as', x')"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   266
  shows "p \<bullet> (supp x \<inter> as) = supp x' \<inter> as'"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   267
proof -
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   268
  from a have "(supp x - as) \<sharp>* p" by  (auto simp only: alphas)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   269
  then have *: "p \<bullet> (supp x - as) = (supp x - as)" 
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   270
    by (simp add: atom_set_perm_eq)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   271
  have "(supp x' - as') \<union> (supp x' \<inter> as') = supp x'" by auto
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   272
  also have "\<dots> = supp (p \<bullet> x)" using a by (simp add: alphas)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   273
  also have "\<dots> = p \<bullet> (supp x)" by (simp add: supp_eqvt)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   274
  also have "\<dots> = p \<bullet> ((supp x - as) \<union> (supp x \<inter> as))" by auto
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   275
  also have "\<dots> = (p \<bullet> (supp x - as)) \<union> (p \<bullet> (supp x \<inter> as))" by (simp add: union_eqvt)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   276
  also have "\<dots> = (supp x - as) \<union> (p \<bullet> (supp x \<inter> as))" using * by simp
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   277
  also have "\<dots> = (supp x' - as') \<union> (p \<bullet> (supp x \<inter> as))" using a by (simp add: alphas)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   278
  finally have "(supp x' - as') \<union> (supp x' \<inter> as') = (supp x' - as') \<union> (p \<bullet> (supp x \<inter> as))" .
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   279
  moreover 
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   280
  have "(supp x' - as') \<inter> (supp x' \<inter> as') = {}" by auto
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   281
  moreover 
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   282
  have "(supp x - as) \<inter> (supp x \<inter> as) = {}" by auto
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   283
  then have "p \<bullet> ((supp x - as) \<inter> (supp x \<inter> as) = {})" by (simp add: permute_bool_def)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   284
  then have "(p \<bullet> (supp x - as)) \<inter> (p \<bullet> (supp x \<inter> as)) = {}" by (perm_simp) (simp)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   285
  then have "(supp x - as) \<inter> (p \<bullet> (supp x \<inter> as)) = {}" using * by simp
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   286
  then have "(supp x' - as') \<inter> (p \<bullet> (supp x \<inter> as)) = {}" using a by (simp add: alphas)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   287
  ultimately show "p \<bullet> (supp x \<inter> as) = supp x' \<inter> as'"
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   288
    by (auto dest: disjoint_right_eq)
2712
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   289
qed
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   290
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   291
lemma alpha_abs_res_stronger1_aux:
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   292
  assumes asm: "(as, x) \<approx>res (op =) supp p' (as', x')"
2669
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   293
  shows "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   294
proof -
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   295
  from asm have 0: "(supp x - as) \<sharp>* p'" by  (auto simp only: alphas)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   296
  then have #: "p' \<bullet> (supp x - as) = (supp x - as)" 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   297
    by (simp add: atom_set_perm_eq)
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   298
  obtain p where *: "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" and **: "supp p \<subseteq> supp x \<union> p' \<bullet> supp x"
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   299
    using set_renaming_perm2 by blast
2669
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   300
  from * have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   301
  from 0 have 1: "(supp x - as) \<sharp>* p" using *
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   302
    by (auto simp add: fresh_star_def fresh_perm)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   303
  then have 2: "(supp x - as) \<inter> supp p = {}"
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   304
    by (auto simp add: fresh_star_def fresh_def)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   305
  have b: "supp x = (supp x - as) \<union> (supp x \<inter> as)" by auto
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   306
  have "supp p \<subseteq> supp x \<union> p' \<bullet> supp x" using ** by simp
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   307
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (p' \<bullet> ((supp x - as) \<union> (supp x \<inter> as)))" 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   308
    using b by simp
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   309
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> ((p' \<bullet> (supp x - as)) \<union> (p' \<bullet> (supp x \<inter> as)))"
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   310
    by (simp add: union_eqvt)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   311
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (p' \<bullet> (supp x \<inter> as))" 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   312
    using # by auto
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   313
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (supp x' \<inter> as')" using asm
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   314
    by (simp add: supp_property_res)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   315
  finally have "supp p \<subseteq> (supp x - as) \<union> (supp x \<inter> as) \<union> (supp x' \<inter> as')" .
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   316
  then 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   317
  have "supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" using 2 by auto
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   318
  moreover 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   319
  have "(as, x) \<approx>res (op =) supp p (as', x')" using asm 1 a by (simp add: alphas)
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   320
  ultimately 
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   321
  show "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" by blast
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   322
qed
1d1772a89026 the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents: 2668
diff changeset
   323
2712
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   324
lemma alpha_abs_res_minimal:
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   325
  assumes asm: "(as, x) \<approx>res (op =) supp p (as', x')"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   326
  shows "(as \<inter> supp x, x) \<approx>res (op =) supp p (as' \<inter> supp x', x')"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   327
  using asm unfolding alpha_res by (auto simp add: Diff_Int)
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   328
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   329
lemma alpha_abs_res_abs_set:
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   330
  assumes asm: "(as, x) \<approx>res (op =) supp p (as', x')"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   331
  shows "(as \<inter> supp x, x) \<approx>set (op =) supp p (as' \<inter> supp x', x')"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   332
proof -
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   333
  have c: "p \<bullet> x = x'"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   334
    using alpha_abs_res_minimal[OF asm] unfolding alpha_res by clarify
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   335
  then have a: "supp x - as \<inter> supp x = supp (p \<bullet> x) - as' \<inter> supp (p \<bullet> x)"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   336
    using alpha_abs_res_minimal[OF asm] by (simp add: alpha_res)
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   337
  have b: "(supp x - as \<inter> supp x) \<sharp>* p"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   338
    using alpha_abs_res_minimal[OF asm] unfolding alpha_res by clarify
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   339
  have "p \<bullet> (as \<inter> supp x) = as' \<inter> supp (p \<bullet> x)"
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   340
    by (metis Int_commute asm c supp_property_res)
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   341
  then show ?thesis using a b c unfolding alpha_set by simp
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   342
qed
c66d288b9fa0 alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2683
diff changeset
   343
2713
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   344
lemma alpha_abs_set_abs_res:
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   345
  assumes asm: "(as \<inter> supp x, x) \<approx>set (op =) supp p (as' \<inter> supp x', x')"
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   346
  shows "(as, x) \<approx>res (op =) supp p (as', x')"
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   347
  using asm unfolding alphas by (auto simp add: Diff_Int)
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   348
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   349
lemma alpha_abs_res_stronger1:
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   350
  assumes asm: "(as, x) \<approx>res (op =) supp p' (as', x')"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   351
  shows "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   352
using alpha_abs_res_stronger1_aux[OF asm] by auto
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   353
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   354
lemma alpha_abs_set_stronger1:
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   355
  assumes asm: "(as, x) \<approx>set (op =) supp p' (as', x')"
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   356
  shows "\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'"
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   357
proof -
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   358
  from asm have 0: "(supp x - as) \<sharp>* p'" by  (auto simp only: alphas)
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   359
  then have #: "p' \<bullet> (supp x - as) = (supp x - as)" 
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   360
    by (simp add: atom_set_perm_eq)
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   361
  obtain p where *: "\<forall>b \<in> (supp x \<union> as). p \<bullet> b = p' \<bullet> b" 
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   362
    and **: "supp p \<subseteq> (supp x \<union> as) \<union> p' \<bullet> (supp x \<union> as)"
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   363
    using set_renaming_perm2 by blast
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   364
  from * have "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" by blast
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   365
  then have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   366
  from * have "\<forall>b \<in> as. p \<bullet> b = p' \<bullet> b" by blast
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   367
  then have zb: "p \<bullet> as = p' \<bullet> as" 
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3060
diff changeset
   368
    apply(auto simp add: permute_set_def)
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   369
    apply(rule_tac x="xa" in exI)
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   370
    apply(simp)
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   371
    done
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   372
  have zc: "p' \<bullet> as = as'" using asm by (simp add: alphas)
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   373
  from 0 have 1: "(supp x - as) \<sharp>* p" using *
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   374
    by (auto simp add: fresh_star_def fresh_perm)
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   375
  then have 2: "(supp x - as) \<inter> supp p = {}"
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   376
    by (auto simp add: fresh_star_def fresh_def)
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   377
  have b: "supp x = (supp x - as) \<union> (supp x \<inter> as)" by auto
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   378
  have "supp p \<subseteq> supp x \<union> as \<union> p' \<bullet> supp x \<union> p' \<bullet> as" using ** using union_eqvt by blast
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   379
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> (p' \<bullet> ((supp x - as) \<union> (supp x \<inter> as))) \<union> p' \<bullet> as" 
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   380
    using b by simp
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   381
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> 
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   382
    ((p' \<bullet> (supp x - as)) \<union> (p' \<bullet> (supp x \<inter> as))) \<union> p' \<bullet> as" by (simp add: union_eqvt)
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   383
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> (p' \<bullet> (supp x \<inter> as)) \<union> p' \<bullet> as"
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   384
    using # by auto
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   385
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> p' \<bullet> ((supp x \<inter> as) \<union> as)" using union_eqvt
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   386
    by auto
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   387
  also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> p' \<bullet> as" 
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   388
    by (metis Int_commute Un_commute sup_inf_absorb)
2673
87ebc706df67 made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents: 2671
diff changeset
   389
  also have "\<dots> = (supp x - as) \<union> as \<union> p' \<bullet> as" by blast
2671
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   390
  finally have "supp p \<subseteq> (supp x - as) \<union> as \<union> p' \<bullet> as" .
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   391
  then have "supp p \<subseteq> as \<union> p' \<bullet> as" using 2 by blast
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   392
  moreover 
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   393
  have "(as, x) \<approx>set (op =) supp p (as', x')" using asm 1 a zb by (simp add: alphas)
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   394
  ultimately 
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   395
  show "\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'" using zc by blast
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   396
qed
eef49daac6c8 alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2669
diff changeset
   397
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   398
lemma alpha_abs_lst_stronger1:
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   399
  assumes asm: "(as, x) \<approx>lst (op =) supp p' (as', x')"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   400
  shows "\<exists>p. (as, x) \<approx>lst (op =) supp p (as', x') \<and> supp p \<subseteq> set as \<union> set as'"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   401
proof -
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   402
  from asm have 0: "(supp x - set as) \<sharp>* p'" by  (auto simp only: alphas)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   403
  then have #: "p' \<bullet> (supp x - set as) = (supp x - set as)" 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   404
    by (simp add: atom_set_perm_eq)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   405
  obtain p where *: "\<forall>b \<in> (supp x \<union> set as). p \<bullet> b = p' \<bullet> b" 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   406
    and **: "supp p \<subseteq> (supp x \<union> set as) \<union> p' \<bullet> (supp x \<union> set as)"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   407
    using set_renaming_perm2 by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   408
  from * have "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   409
  then have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   410
  from * have "\<forall>b \<in> set as. p \<bullet> b = p' \<bullet> b" by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   411
  then have zb: "p \<bullet> as = p' \<bullet> as" by (induct as) (auto)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   412
  have zc: "p' \<bullet> set as = set as'" using asm by (simp add: alphas set_eqvt)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   413
  from 0 have 1: "(supp x - set as) \<sharp>* p" using *
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   414
    by (auto simp add: fresh_star_def fresh_perm)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   415
  then have 2: "(supp x - set as) \<inter> supp p = {}"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   416
    by (auto simp add: fresh_star_def fresh_def)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   417
  have b: "supp x = (supp x - set as) \<union> (supp x \<inter> set as)" by auto
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   418
  have "supp p \<subseteq> supp x \<union> set as \<union> p' \<bullet> supp x \<union> p' \<bullet> set as" using ** using union_eqvt by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   419
  also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   420
    (p' \<bullet> ((supp x - set as) \<union> (supp x \<inter> set as))) \<union> p' \<bullet> set as" using b by simp
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   421
  also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   422
    ((p' \<bullet> (supp x - set as)) \<union> (p' \<bullet> (supp x \<inter> set as))) \<union> p' \<bullet> set as" by (simp add: union_eqvt)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   423
  also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   424
    (p' \<bullet> (supp x \<inter> set as)) \<union> p' \<bullet> set as" using # by auto
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   425
  also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> p' \<bullet> ((supp x \<inter> set as) \<union> set as)" 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   426
    using union_eqvt by auto
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   427
  also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> p' \<bullet> set as" 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   428
    by (metis Int_commute Un_commute sup_inf_absorb)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   429
  also have "\<dots> = (supp x - set as) \<union> set as \<union> p' \<bullet> set as" by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   430
  finally have "supp p \<subseteq> (supp x - set as) \<union> set as \<union> p' \<bullet> set as" .
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   431
  then have "supp p \<subseteq> set as \<union> p' \<bullet> set as" using 2 by blast
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   432
  moreover 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   433
  have "(as, x) \<approx>lst (op =) supp p (as', x')" using asm 1 a zb by (simp add: alphas)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   434
  ultimately 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   435
  show "\<exists>p. (as, x) \<approx>lst (op =) supp p (as', x') \<and> supp p \<subseteq> set as \<union> set as'" using zc by blast 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   436
qed
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   437
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   438
lemma alphas_abs_stronger:
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   439
  shows "(as, x) \<approx>abs_set (as', x') \<longleftrightarrow> (\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as')"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   440
  and   "(as, x) \<approx>abs_res (as', x') \<longleftrightarrow> (\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as')"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   441
  and   "(bs, x) \<approx>abs_lst (bs', x') \<longleftrightarrow> 
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   442
   (\<exists>p. (bs, x) \<approx>lst (op =) supp p (bs', x') \<and> supp p \<subseteq> set bs \<union> set bs')"
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   443
apply(rule iffI)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   444
apply(auto simp add: alphas_abs alpha_abs_set_stronger1)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   445
apply(auto simp add: alphas_abs)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   446
apply(rule iffI)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   447
apply(auto simp add: alphas_abs alpha_abs_res_stronger1)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   448
apply(auto simp add: alphas_abs)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   449
apply(rule iffI)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   450
apply(auto simp add: alphas_abs alpha_abs_lst_stronger1)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   451
apply(auto simp add: alphas_abs)[1]
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   452
done
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   453
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   454
lemma alpha_res_alpha_set:
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   455
  "(bs, x) \<approx>res op = supp p (cs, y) \<longleftrightarrow> (bs \<inter> supp x, x) \<approx>set op = supp p (cs \<inter> supp y, y)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   456
  using alpha_abs_set_abs_res alpha_abs_res_abs_set by blast
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   457
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   458
section {* Quotient types *}
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   459
3172
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   460
(* FIXME: The three could be defined together, but due to bugs
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   461
   introduced by the lifting package it doesn't work anymore *)
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   462
quotient_type
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   463
    'a abs_set = "(atom set \<times> 'a::pt)" / "alpha_abs_set"
3172
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   464
  apply(rule_tac [!] equivpI)
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   465
  unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   466
  by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:)
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   467
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   468
quotient_type
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   469
    'b abs_res = "(atom set \<times> 'b::pt)" / "alpha_abs_res"
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   470
  apply(rule_tac [!] equivpI)
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   471
  unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   472
  by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:)
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   473
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   474
quotient_type
4cf3a4d36799 Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents: 3157
diff changeset
   475
   'c abs_lst = "(atom list \<times> 'c::pt)" / "alpha_abs_lst"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   476
  apply(rule_tac [!] equivpI)
2592
98236fbd8aa6 updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents: 2591
diff changeset
   477
  unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   478
  by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
quotient_definition
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   481
  Abs_set ("[_]set. _" [60, 60] 60)
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   482
where
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   483
  "Abs_set::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_set"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
is
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   485
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" .
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   487
quotient_definition
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   488
  Abs_res ("[_]res. _" [60, 60] 60)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   489
where
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   490
  "Abs_res::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_res"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   491
is
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   492
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" .
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   493
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   494
quotient_definition
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   495
  Abs_lst ("[_]lst. _" [60, 60] 60)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   496
where
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   497
  "Abs_lst::atom list \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_lst"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   498
is
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   499
  "Pair::atom list \<Rightarrow> ('a::pt) \<Rightarrow> (atom list \<times> 'a)" .
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   500
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
lemma [quot_respect]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   502
  shows "(op= ===> op= ===> alpha_abs_set) Pair Pair"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   503
  and   "(op= ===> op= ===> alpha_abs_res) Pair Pair"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   504
  and   "(op= ===> op= ===> alpha_abs_lst) Pair Pair"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   505
  unfolding fun_rel_def
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   506
  by (auto intro: alphas_abs_refl)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
lemma [quot_respect]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   509
  shows "(op= ===> alpha_abs_set ===> alpha_abs_set) permute permute"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   510
  and   "(op= ===> alpha_abs_res ===> alpha_abs_res) permute permute"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   511
  and   "(op= ===> alpha_abs_lst ===> alpha_abs_lst) permute permute"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   512
  unfolding fun_rel_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   513
  by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   515
lemma Abs_eq_iff:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   516
  shows "[bs]set. x = [bs']set. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op =) supp p (bs', y))"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   517
  and   "[bs]res. x = [bs']res. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (bs', y))"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   518
  and   "[cs]lst. x = [cs']lst. y \<longleftrightarrow> (\<exists>p. (cs, x) \<approx>lst (op =) supp p (cs', y))"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   519
  by (lifting alphas_abs)
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   520
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   521
lemma Abs_eq_iff2:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   522
  shows "[bs]set. x = [bs']set. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op=) supp p (bs', y) \<and> supp p \<subseteq> bs \<union> bs')"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   523
  and   "[bs]res. x = [bs']res. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op=) supp p (bs', y) \<and> supp p \<subseteq> bs \<union> bs')"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   524
  and   "[cs]lst. x = [cs']lst. y \<longleftrightarrow> (\<exists>p. (cs, x) \<approx>lst (op=) supp p (cs', y) \<and> supp p \<subseteq> set cs \<union> set cs')"
2674
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   525
  by (lifting alphas_abs_stronger)
3c79dfec1cf0 derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents: 2673
diff changeset
   526
3024
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3004
diff changeset
   527
2713
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   528
lemma Abs_eq_res_set:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   529
  shows "[bs]res. x = [cs]res. y \<longleftrightarrow> [bs \<inter> supp x]set. x = [cs \<inter> supp y]set. y"
3024
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3004
diff changeset
   530
  unfolding Abs_eq_iff alpha_res_alpha_set by rule
2713
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   531
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   532
lemma Abs_eq_res_supp:
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   533
  assumes asm: "supp x \<subseteq> bs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   534
  shows "[as]res. x = [as \<inter> bs]res. x"
2713
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   535
  unfolding Abs_eq_iff alphas
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   536
  apply (rule_tac x="0::perm" in exI)
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   537
  apply (simp add: fresh_star_zero)
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   538
  using asm by blast
a84999edbcb3 More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2712
diff changeset
   539
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   540
lemma Abs_exhausts:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   541
  shows "(\<And>as (x::'a::pt). y1 = [as]set. x \<Longrightarrow> P1) \<Longrightarrow> P1"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   542
  and   "(\<And>as (x::'a::pt). y2 = [as]res. x \<Longrightarrow> P2) \<Longrightarrow> P2"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   543
  and   "(\<And>bs (x::'a::pt). y3 = [bs]lst. x \<Longrightarrow> P3) \<Longrightarrow> P3"
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   544
  by (lifting prod.exhaust[where 'a="atom set" and 'b="'a"]
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   545
              prod.exhaust[where 'a="atom set" and 'b="'a"]
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   546
              prod.exhaust[where 'a="atom list" and 'b="'a"])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   548
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   549
instantiation abs_set :: (pt) pt
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
quotient_definition
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   553
  "permute_abs_set::perm \<Rightarrow> ('a::pt abs_set) \<Rightarrow> 'a abs_set"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
is
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   555
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)" 
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   556
  by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   558
lemma permute_Abs_set[simp]:
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   559
  fixes x::"'a::pt"  
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   560
  shows "(p \<bullet> ([as]set. x)) = [p \<bullet> as]set. (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   561
  by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
instance
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   565
  apply(case_tac [!] x rule: Abs_exhausts(1))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   566
  apply(simp_all)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   567
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   568
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   569
end
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   570
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   571
instantiation abs_res :: (pt) pt
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   572
begin
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   573
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   574
quotient_definition
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   575
  "permute_abs_res::perm \<Rightarrow> ('a::pt abs_res) \<Rightarrow> 'a abs_res"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   576
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   577
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   578
  by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   579
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   580
lemma permute_Abs_res[simp]:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   581
  fixes x::"'a::pt"  
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   582
  shows "(p \<bullet> ([as]res. x)) = [p \<bullet> as]res. (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   583
  by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"])
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   584
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   585
instance
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   586
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   587
  apply(case_tac [!] x rule: Abs_exhausts(2))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   588
  apply(simp_all)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   589
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   590
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   591
end
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   592
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   593
instantiation abs_lst :: (pt) pt
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   594
begin
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   595
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   596
quotient_definition
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   597
  "permute_abs_lst::perm \<Rightarrow> ('a::pt abs_lst) \<Rightarrow> 'a abs_lst"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   598
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   599
  "permute:: perm \<Rightarrow> (atom list \<times> 'a::pt) \<Rightarrow> (atom list \<times> 'a::pt)"
3152
da59c94bed7e updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   600
  by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   601
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   602
lemma permute_Abs_lst[simp]:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   603
  fixes x::"'a::pt"  
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   604
  shows "(p \<bullet> ([as]lst. x)) = [p \<bullet> as]lst. (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   605
  by (lifting permute_prod.simps[where 'a="atom list" and 'b="'a"])
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   606
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   607
instance
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   608
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   609
  apply(case_tac [!] x rule: Abs_exhausts(3))
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   615
lemmas permute_Abs[eqvt] = permute_Abs_set permute_Abs_res permute_Abs_lst
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   616
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   617
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   618
lemma Abs_swap1:
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   619
  assumes a1: "a \<notin> (supp x) - bs"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   620
  and     a2: "b \<notin> (supp x) - bs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   621
  shows "[bs]set. x = [(a \<rightleftharpoons> b) \<bullet> bs]set. ((a \<rightleftharpoons> b) \<bullet> x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   622
  and   "[bs]res. x = [(a \<rightleftharpoons> b) \<bullet> bs]res. ((a \<rightleftharpoons> b) \<bullet> x)"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   623
  unfolding Abs_eq_iff
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   624
  unfolding alphas
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   625
  unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] 
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   626
  unfolding fresh_star_def fresh_def
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   627
  unfolding swap_set_not_in[OF a1 a2] 
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   628
  using a1 a2
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   629
  by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   630
     (auto simp add: supp_perm swap_atom)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   631
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   632
lemma Abs_swap2:
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   633
  assumes a1: "a \<notin> (supp x) - (set bs)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   634
  and     a2: "b \<notin> (supp x) - (set bs)"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   635
  shows "[bs]lst. x = [(a \<rightleftharpoons> b) \<bullet> bs]lst. ((a \<rightleftharpoons> b) \<bullet> x)"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   636
  unfolding Abs_eq_iff
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   637
  unfolding alphas
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   638
  unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] set_eqvt[symmetric]
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   639
  unfolding fresh_star_def fresh_def
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   640
  unfolding swap_set_not_in[OF a1 a2]
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   641
  using a1 a2
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   642
  by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   643
     (auto simp add: supp_perm swap_atom)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   644
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   645
lemma Abs_supports:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   646
  shows "((supp x) - as) supports ([as]set. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   647
  and   "((supp x) - as) supports ([as]res. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   648
  and   "((supp x) - set bs) supports ([bs]lst. x)"
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   649
  unfolding supports_def
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   650
  unfolding permute_Abs
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   651
  by (simp_all add: Abs_swap1[symmetric] Abs_swap2[symmetric])
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   652
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   653
function
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   654
  supp_set  :: "('a::pt) abs_set \<Rightarrow> atom set"
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   655
where
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   656
  "supp_set ([as]set. x) = supp x - as"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   657
apply(case_tac x rule: Abs_exhausts(1))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   658
apply(simp)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   659
apply(simp add: Abs_eq_iff alphas_abs alphas)
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   660
done
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   661
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   662
termination supp_set 
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   663
  by lexicographic_order
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   665
function
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   666
  supp_res :: "('a::pt) abs_res \<Rightarrow> atom set"
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   667
where
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   668
  "supp_res ([as]res. x) = supp x - as"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   669
apply(case_tac x rule: Abs_exhausts(2))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   670
apply(simp)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   671
apply(simp add: Abs_eq_iff alphas_abs alphas)
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   672
done
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   673
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   674
termination supp_res 
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   675
  by lexicographic_order
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   677
function
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   678
  supp_lst :: "('a::pt) abs_lst \<Rightarrow> atom set"
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   679
where
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   680
  "supp_lst (Abs_lst cs x) = (supp x) - (set cs)"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   681
apply(case_tac x rule: Abs_exhausts(3))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   682
apply(simp)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   683
apply(simp add: Abs_eq_iff alphas_abs alphas)
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   684
done
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   686
termination supp_lst
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   687
  by lexicographic_order
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   688
2663
54aade5d0fe6 moved high level code from LamTest into the main libraries.
Christian Urban <urbanc@in.tum.de>
parents: 2659
diff changeset
   689
lemma supp_funs_eqvt[eqvt]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   690
  shows "(p \<bullet> supp_set x) = supp_set (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   691
  and   "(p \<bullet> supp_res y) = supp_res (p \<bullet> y)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   692
  and   "(p \<bullet> supp_lst z) = supp_lst (p \<bullet> z)"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   693
  apply(case_tac x rule: Abs_exhausts(1))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   694
  apply(simp add: supp_eqvt Diff_eqvt)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   695
  apply(case_tac y rule: Abs_exhausts(2))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   696
  apply(simp add: supp_eqvt Diff_eqvt)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   697
  apply(case_tac z rule: Abs_exhausts(3))
1686
7b3dd407f6b3 got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents: 1673
diff changeset
   698
  apply(simp add: supp_eqvt Diff_eqvt set_eqvt)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   701
lemma Abs_fresh_aux:
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   702
  shows "a \<sharp> [bs]set. x \<Longrightarrow> a \<sharp> supp_set ([bs]set. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   703
  and   "a \<sharp> [bs]res. x \<Longrightarrow> a \<sharp> supp_res ([bs]res. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   704
  and   "a \<sharp> [cs]lst. x \<Longrightarrow> a \<sharp> supp_lst ([cs]lst. x)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   705
  by (rule_tac [!] fresh_fun_eqvt_app)
2663
54aade5d0fe6 moved high level code from LamTest into the main libraries.
Christian Urban <urbanc@in.tum.de>
parents: 2659
diff changeset
   706
     (auto simp only: eqvt_def eqvts_raw)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   707
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   708
lemma Abs_supp_subset1:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   709
  assumes a: "finite (supp x)"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   710
  shows "(supp x) - as \<subseteq> supp ([as]set. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   711
  and   "(supp x) - as \<subseteq> supp ([as]res. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   712
  and   "(supp x) - (set bs) \<subseteq> supp ([bs]lst. x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   713
  unfolding supp_conv_fresh
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   714
  by (auto dest!: Abs_fresh_aux)
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   715
     (simp_all add: fresh_def supp_finite_atom_set a)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   717
lemma Abs_supp_subset2:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   718
  assumes a: "finite (supp x)"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   719
  shows "supp ([as]set. x) \<subseteq> (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   720
  and   "supp ([as]res. x) \<subseteq> (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   721
  and   "supp ([bs]lst. x) \<subseteq> (supp x) - (set bs)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   722
  by (rule_tac [!] supp_is_subset)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   723
     (simp_all add: Abs_supports a)
1478
1ea4ca823266 added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 1470
diff changeset
   724
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   725
lemma Abs_finite_supp:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   726
  assumes a: "finite (supp x)"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   727
  shows "supp ([as]set. x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   728
  and   "supp ([as]res. x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   729
  and   "supp ([bs]lst. x) = (supp x) - (set bs)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   730
using Abs_supp_subset1[OF a] Abs_supp_subset2[OF a]
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   731
  by blast+
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   733
lemma supp_Abs:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
  fixes x::"'a::fs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   735
  shows "supp ([as]set. x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   736
  and   "supp ([as]res. x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   737
  and   "supp ([bs]lst. x) = (supp x) - (set bs)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   738
by (simp_all add: Abs_finite_supp finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   740
instance abs_set :: (fs) fs
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   742
  apply(case_tac x rule: Abs_exhausts(1))
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   743
  apply(simp add: supp_Abs finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   746
instance abs_res :: (fs) fs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   747
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   748
  apply(case_tac x rule: Abs_exhausts(2))
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   749
  apply(simp add: supp_Abs finite_supp)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   750
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   751
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   752
instance abs_lst :: (fs) fs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   753
  apply(default)
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   754
  apply(case_tac x rule: Abs_exhausts(3))
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   755
  apply(simp add: supp_Abs finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   758
lemma Abs_fresh_iff:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   759
  fixes x::"'a::fs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   760
  shows "a \<sharp> [bs]set. x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   761
  and   "a \<sharp> [bs]res. x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   762
  and   "a \<sharp> [cs]lst. x \<longleftrightarrow> a \<in> (set cs) \<or> (a \<notin> (set cs) \<and> a \<sharp> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   763
  unfolding fresh_def
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   764
  unfolding supp_Abs
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   765
  by auto
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   766
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   767
lemma Abs_fresh_star_iff:
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   768
  fixes x::"'a::fs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   769
  shows "as \<sharp>* ([bs]set. x) \<longleftrightarrow> (as - bs) \<sharp>* x"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   770
  and   "as \<sharp>* ([bs]res. x) \<longleftrightarrow> (as - bs) \<sharp>* x"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   771
  and   "as \<sharp>* ([cs]lst. x) \<longleftrightarrow> (as - set cs) \<sharp>* x"
2591
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   772
  unfolding fresh_star_def
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   773
  by (auto simp add: Abs_fresh_iff)
35c570891a3a isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents: 2584
diff changeset
   774
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   775
lemma Abs_fresh_star:
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   776
  fixes x::"'a::fs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   777
  shows "as \<subseteq> as' \<Longrightarrow> as \<sharp>* ([as']set. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   778
  and   "as \<subseteq> as' \<Longrightarrow> as \<sharp>* ([as']res. x)"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   779
  and   "bs \<subseteq> set bs' \<Longrightarrow> bs \<sharp>* ([bs']lst. x)"
2491
d0961e6d6881 more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2489
diff changeset
   780
  unfolding fresh_star_def
2584
1eac050a36f4 completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents: 2574
diff changeset
   781
  by(auto simp add: Abs_fresh_iff)
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   782
2730
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   783
lemma Abs_fresh_star2:
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   784
  fixes x::"'a::fs"
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   785
  shows "as \<inter> bs = {} \<Longrightarrow> as \<sharp>* ([bs]set. x) \<longleftrightarrow> as \<sharp>* x"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   786
  and   "as \<inter> bs = {} \<Longrightarrow> as \<sharp>* ([bs]res. x) \<longleftrightarrow> as \<sharp>* x"
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   787
  and   "cs \<inter> set ds = {} \<Longrightarrow> cs \<sharp>* ([ds]lst. x) \<longleftrightarrow> cs \<sharp>* x"
2730
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   788
  unfolding fresh_star_def Abs_fresh_iff
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   789
  by auto
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   790
eebc24b9cf39 added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents: 2713
diff changeset
   791
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   792
section {* Abstractions of single atoms *}
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   793
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   794
lemma Abs1_eq:
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   795
  fixes x y::"'a::fs"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   796
  shows "[{atom a}]set. x = [{atom a}]set. y \<longleftrightarrow> x = y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   797
  and   "[{atom a}]res. x = [{atom a}]res. y \<longleftrightarrow> x = y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   798
  and   "[[atom a]]lst. x = [[atom a]]lst. y \<longleftrightarrow> x = y"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   799
unfolding Abs_eq_iff2 alphas
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   800
apply(simp_all add: supp_perm_singleton fresh_star_def fresh_zero_perm)
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   801
apply(blast)+
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   802
done
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   803
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   804
lemma Abs1_eq_fresh:
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   805
  fixes x y::"'a::fs"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   806
  and a b c::"'b::at"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   807
  assumes "atom c \<sharp> (a, b, x, y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   808
  shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   809
  and   "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   810
  and   "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   811
proof -
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   812
  have "[{atom a}]set. x = (a \<leftrightarrow> c) \<bullet> ([{atom a}]set. x)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   813
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   814
  then have "[{atom a}]set. x = [{atom c}]set. ((a \<leftrightarrow> c) \<bullet> x)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   815
  moreover 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   816
  have "[{atom b}]set. y = (b \<leftrightarrow> c) \<bullet> ([{atom b}]set. y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   817
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   818
  then have "[{atom b}]set. y = [{atom c}]set. ((b \<leftrightarrow> c) \<bullet> y)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   819
  ultimately 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   820
  show "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   821
    by (simp add: Abs1_eq)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   822
next
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   823
  have "[{atom a}]res. x = (a \<leftrightarrow> c) \<bullet> ([{atom a}]res. x)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   824
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   825
  then have "[{atom a}]res. x = [{atom c}]res. ((a \<leftrightarrow> c) \<bullet> x)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   826
  moreover 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   827
  have "[{atom b}]res. y = (b \<leftrightarrow> c) \<bullet> ([{atom b}]res. y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   828
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   829
  then have "[{atom b}]res. y = [{atom c}]res. ((b \<leftrightarrow> c) \<bullet> y)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   830
  ultimately 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   831
  show "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   832
    by (simp add: Abs1_eq)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   833
next
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   834
  have "[[atom a]]lst. x = (a \<leftrightarrow> c) \<bullet> ([[atom a]]lst. x)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   835
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   836
  then have "[[atom a]]lst. x = [[atom c]]lst. ((a \<leftrightarrow> c) \<bullet> x)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   837
  moreover 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   838
  have "[[atom b]]lst. y = (b \<leftrightarrow> c) \<bullet> ([[atom b]]lst. y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   839
    by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   840
  then have "[[atom b]]lst. y = [[atom c]]lst. ((b \<leftrightarrow> c) \<bullet> y)" by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   841
  ultimately 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   842
  show "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   843
    by (simp add: Abs1_eq)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   844
qed
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   845
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   846
lemma Abs1_eq_all:
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   847
  fixes x y::"'a::fs"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   848
  and z::"'c::fs"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   849
  and a b::"'b::at"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   850
  shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   851
  and   "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   852
  and   "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   853
apply -
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   854
apply(auto) 
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   855
apply(simp add: Abs1_eq_fresh(1)[symmetric])
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   856
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   857
apply(drule_tac x="aa" in spec)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   858
apply(simp)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   859
apply(subst Abs1_eq_fresh(1))
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   860
apply(auto simp add: fresh_Pair)[1]
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   861
apply(assumption)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   862
apply(simp add: Abs1_eq_fresh(2)[symmetric])
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   863
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   864
apply(drule_tac x="aa" in spec)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   865
apply(simp)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   866
apply(subst Abs1_eq_fresh(2))
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   867
apply(auto simp add: fresh_Pair)[1]
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   868
apply(assumption)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   869
apply(simp add: Abs1_eq_fresh(3)[symmetric])
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   870
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   871
apply(drule_tac x="aa" in spec)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   872
apply(simp)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   873
apply(subst Abs1_eq_fresh(3))
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   874
apply(auto simp add: fresh_Pair)[1]
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   875
apply(assumption)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   876
done
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   877
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   878
lemma Abs1_eq_iff:
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   879
  fixes x y::"'a::fs"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   880
  and a b::"'b::at"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   881
  shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   882
  and   "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   883
  and   "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   884
proof -
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   885
  { assume "a = b"
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   886
    then have "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   887
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   888
  moreover
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   889
  { assume *: "a \<noteq> b" and **: "Abs_set {atom a} x = Abs_set {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   890
    have #: "atom a \<sharp> Abs_set {atom b} y" by (simp add: **[symmetric] Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   891
    have "Abs_set {atom a} ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> (Abs_set {atom b} y)" by (simp)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   892
    also have "\<dots> = Abs_set {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   893
      by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   894
    also have "\<dots> = Abs_set {atom a} x" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   895
    finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   896
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   897
  moreover 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   898
  { assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   899
    have "Abs_set {atom a} x = Abs_set {atom a} ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   900
    also have "\<dots> = (a \<leftrightarrow> b) \<bullet> Abs_set {atom b} y" by (simp add: permute_set_def assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   901
    also have "\<dots> = Abs_set {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   902
      by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   903
    finally have "Abs_set {atom a} x = Abs_set {atom b} y" .
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   904
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   905
  ultimately 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   906
  show "Abs_set {atom a} x = Abs_set {atom b} y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   907
    by blast
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   908
next
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   909
  { assume "a = b"
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   910
    then have "Abs_res {atom a} x = Abs_res {atom b} y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   911
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   912
  moreover
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   913
  { assume *: "a \<noteq> b" and **: "Abs_res {atom a} x = Abs_res {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   914
    have #: "atom a \<sharp> Abs_res {atom b} y" by (simp add: **[symmetric] Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   915
    have "Abs_res {atom a} ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> (Abs_res {atom b} y)" by (simp add: assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   916
    also have "\<dots> = Abs_res {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   917
      by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   918
    also have "\<dots> = Abs_res {atom a} x" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   919
    finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   920
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   921
  moreover 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   922
  { assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   923
    have "Abs_res {atom a} x = Abs_res {atom a} ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   924
    also have "\<dots> = (a \<leftrightarrow> b) \<bullet> Abs_res {atom b} y" by (simp add: permute_set_def assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   925
    also have "\<dots> = Abs_res {atom b} y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   926
      by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   927
    finally have "Abs_res {atom a} x = Abs_res {atom b} y" .
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   928
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   929
  ultimately 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   930
  show "Abs_res {atom a} x = Abs_res {atom b} y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   931
    by blast
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   932
next
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   933
  { assume "a = b"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   934
    then have "Abs_lst [atom a] x = Abs_lst [atom b] y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   935
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   936
  moreover
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   937
  { assume *: "a \<noteq> b" and **: "Abs_lst [atom a] x = Abs_lst [atom b] y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   938
    have #: "atom a \<sharp> Abs_lst [atom b] y" by (simp add: **[symmetric] Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   939
    have "Abs_lst [atom a] ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> (Abs_lst [atom b] y)" by (simp add: assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   940
    also have "\<dots> = Abs_lst [atom b] y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   941
      by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   942
    also have "\<dots> = Abs_lst [atom a] x" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   943
    finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff)
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   944
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   945
  moreover 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   946
  { assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   947
    have "Abs_lst [atom a] x = Abs_lst [atom a] ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   948
    also have "\<dots> = (a \<leftrightarrow> b) \<bullet> Abs_lst [atom b] y" by (simp add: assms)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   949
    also have "\<dots> = Abs_lst [atom b] y"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   950
      by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff)
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   951
    finally have "Abs_lst [atom a] x = Abs_lst [atom b] y" .
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   952
  }
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   953
  ultimately 
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   954
  show "Abs_lst [atom a] x = Abs_lst [atom b] y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)"
2679
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   955
    by blast
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   956
qed
e003e5e36bae added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents: 2674
diff changeset
   957
2683
42c0d011a177 ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents: 2679
diff changeset
   958
lemma Abs1_eq_iff':
42c0d011a177 ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents: 2679
diff changeset
   959
  fixes x::"'a::fs"
3191
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   960
  and a b::"'b::at"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   961
  shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   962
  and   "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)"
0440bc1a2438 streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
   963
  and   "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)"
2683
42c0d011a177 ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents: 2679
diff changeset
   964
using assms by (auto simp add: Abs1_eq_iff fresh_permute_left)
42c0d011a177 ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents: 2679
diff changeset
   965
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   966
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   967
subsection {* Renaming of bodies of abstractions *}
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   968
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   969
lemma Abs_rename_set:
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   970
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
   971
  assumes a: "(p \<bullet> bs) \<sharp>* x" 
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   972
  (*and     b: "finite bs"*)
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
   973
  shows "\<exists>q. [bs]set. x = [p \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   974
proof -
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   975
  from set_renaming_perm2 
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   976
  obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   977
  have ***: "q \<bullet> bs = p \<bullet> bs" using *
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   978
    unfolding permute_set_eq_image image_def by auto
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   979
  have "[bs]set. x =  q \<bullet> ([bs]set. x)"
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   980
    apply(rule perm_supp_eq[symmetric])
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   981
    using a **
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   982
    unfolding Abs_fresh_star_iff
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   983
    unfolding fresh_star_def
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   984
    by auto
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   985
  also have "\<dots> = [q \<bullet> bs]set. (q \<bullet> x)" by simp
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   986
  finally have "[bs]set. x = [p \<bullet> bs]set. (q \<bullet> x)" by (simp add: ***)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   987
  then show "\<exists>q. [bs]set. x = [p \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   988
qed
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   989
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   990
lemma Abs_rename_res:
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   991
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
   992
  assumes a: "(p \<bullet> bs) \<sharp>* x" 
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   993
  (*and     b: "finite bs"*)
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
   994
  shows "\<exists>q. [bs]res. x = [p \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
   995
proof -
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
   996
  from set_renaming_perm2 
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
   997
  obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   998
  have ***: "q \<bullet> bs = p \<bullet> bs" using * 
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
   999
    unfolding permute_set_eq_image image_def by auto
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1000
  have "[bs]res. x =  q \<bullet> ([bs]res. x)"
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1001
    apply(rule perm_supp_eq[symmetric])
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1002
    using a **
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1003
    unfolding Abs_fresh_star_iff
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1004
    unfolding fresh_star_def
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1005
    by auto
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1006
  also have "\<dots> = [q \<bullet> bs]res. (q \<bullet> x)" by simp
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1007
  finally have "[bs]res. x = [p \<bullet> bs]res. (q \<bullet> x)" by (simp add: ***)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1008
  then show "\<exists>q. [bs]res. x = [p \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1009
qed
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1010
2611
3d101f2f817c simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents: 2599
diff changeset
  1011
lemma Abs_rename_lst:
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1012
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
  1013
  assumes a: "(p \<bullet> (set bs)) \<sharp>* x" 
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1014
  shows "\<exists>q. [bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1015
proof -
3058
Christian Urban <urbanc@in.tum.de>
parents: 3024
diff changeset
  1016
  from list_renaming_perm 
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1017
  obtain q where *: "\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> set bs \<union> (p \<bullet> set bs)" by blast
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1018
  have ***: "q \<bullet> bs = p \<bullet> bs" using * by (induct bs) (simp_all add: insert_eqvt)
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1019
  have "[bs]lst. x =  q \<bullet> ([bs]lst. x)"
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1020
    apply(rule perm_supp_eq[symmetric])
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1021
    using a **
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1022
    unfolding Abs_fresh_star_iff
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1023
    unfolding fresh_star_def
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1024
    by auto
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1025
  also have "\<dots> = [q \<bullet> bs]lst. (q \<bullet> x)" by simp
2668
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1026
  finally have "[bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x)" by (simp add: ***)
92c001d93225 modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2663
diff changeset
  1027
  then show "\<exists>q. [bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1028
qed
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1029
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1030
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1031
text {* for deep recursive binders *}
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1032
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1033
lemma Abs_rename_set':
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1034
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
  1035
  assumes a: "(p \<bullet> bs) \<sharp>* x" 
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
  1036
  (*and     b: "finite bs"*)
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1037
  shows "\<exists>q. [bs]set. x = [q \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
  1038
using Abs_rename_set[OF a] by metis
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1039
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1040
lemma Abs_rename_res':
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1041
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
  1042
  assumes a: "(p \<bullet> bs) \<sharp>* x" 
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
  1043
  (*and     b: "finite bs"*)
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1044
  shows "\<exists>q. [bs]res. x = [q \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
3060
6613514ff6cb tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents: 3058
diff changeset
  1045
using Abs_rename_res[OF a] by metis
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1046
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1047
lemma Abs_rename_lst':
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1048
  fixes x::"'a::fs"
2659
619ecb57db38 strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2635
diff changeset
  1049
  assumes a: "(p \<bullet> (set bs)) \<sharp>* x" 
2616
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1050
  shows "\<exists>q. [bs]lst. x = [q \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs"
dd7490fdd998 all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
  1051
using Abs_rename_lst[OF a] by metis
2599
d6fe94028a5d moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents: 2592
diff changeset
  1052
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
  1053
section {* Infrastructure for building tuples of relations and functions *}
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
  1054
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1055
fun
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1056
  prod_fv :: "('a \<Rightarrow> atom set) \<Rightarrow> ('b \<Rightarrow> atom set) \<Rightarrow> ('a \<times> 'b) \<Rightarrow> atom set"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1057
where
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1058
  "prod_fv fv1 fv2 (x, y) = fv1 x \<union> fv2 y"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1059
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1060
definition 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1061
  prod_alpha :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool)"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1062
where
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1063
 "prod_alpha = prod_rel"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1064
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1065
lemma [quot_respect]:
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1066
  shows "((R1 ===> op =) ===> (R2 ===> op =) ===> prod_rel R1 R2 ===> op =) prod_fv prod_fv"
2559
add799cf0817 adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents: 2491
diff changeset
  1067
  unfolding fun_rel_def
add799cf0817 adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents: 2491
diff changeset
  1068
  unfolding prod_rel_def
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1069
  by auto
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1070
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1071
lemma [quot_preserve]:
3157
de89c95c5377 updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents: 3152
diff changeset
  1072
  assumes q1: "Quotient3 R1 abs1 rep1"
de89c95c5377 updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents: 3152
diff changeset
  1073
  and     q2: "Quotient3 R2 abs2 rep2"
2574
50da1be9a7e5 current isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2568
diff changeset
  1074
  shows "((abs1 ---> id) ---> (abs2 ---> id) ---> map_pair rep1 rep2 ---> id) prod_fv = prod_fv"
3157
de89c95c5377 updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents: 3152
diff changeset
  1075
  by (simp add: fun_eq_iff Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1076
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1077
lemma [mono]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1078
  shows "A <= B \<Longrightarrow> C <= D ==> prod_alpha A C <= prod_alpha B D"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1079
  unfolding prod_alpha_def
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1080
  by auto
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1081
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1082
lemma [eqvt]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1083
  shows "p \<bullet> prod_alpha A B x y = prod_alpha (p \<bullet> A) (p \<bullet> B) (p \<bullet> x) (p \<bullet> y)"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1084
  unfolding prod_alpha_def
2559
add799cf0817 adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents: 2491
diff changeset
  1085
  unfolding prod_rel_def
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1086
  by (perm_simp) (rule refl)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1087
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1088
lemma [eqvt]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1089
  shows "p \<bullet> prod_fv A B (x, y) = prod_fv (p \<bullet> A) (p \<bullet> B) (p \<bullet> x, p \<bullet> y)"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1090
  unfolding prod_fv.simps
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1091
  by (perm_simp) (rule refl)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1092
2489
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1093
lemma prod_fv_supp:
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1094
  shows "prod_fv supp supp = supp"
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1095
by (rule ext)
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1096
   (auto simp add: prod_fv.simps supp_Pair)
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1097
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1098
lemma prod_alpha_eq:
c0695bb33fcd added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents: 2479
diff changeset
  1099
  shows "prod_alpha (op=) (op=) = (op=)"
2843
1ae3c9b2d557 Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2733
diff changeset
  1100
  unfolding prod_alpha_def
1ae3c9b2d557 Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2733
diff changeset
  1101
  by (auto intro!: ext)
1ae3c9b2d557 Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2733
diff changeset
  1102
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
  1103
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105