Nominal/Abs.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 04 Sep 2010 07:39:38 +0800
changeset 2471 894599a50af3
parent 2469 4a6e78bd9de9
child 2473 a3711f07449b
permissions -rw-r--r--
got rid of Nominal2_Supp (is now in Nomina2_Base)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Abs
2467
67b3933c3190 got rid of Nominal_Atoms (folded into Nominal2_Base)
Christian Urban <urbanc@in.tum.de>
parents: 2460
diff changeset
     2
imports "../Nominal-General/Nominal2_Base" 
1804
81b171e2d6d5 added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents: 1744
diff changeset
     3
        "../Nominal-General/Nominal2_Eqvt" 
81b171e2d6d5 added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents: 1744
diff changeset
     4
        "Quotient" 
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1932
diff changeset
     5
        "Quotient_List"
1804
81b171e2d6d5 added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents: 1744
diff changeset
     6
        "Quotient_Product" 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
fun
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    10
  alpha_set 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
where
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[simp del]:
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    13
  "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
    14
     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
    15
     (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
    16
     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
    17
     pi \<bullet> bs = cs"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    19
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    20
  alpha_res
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    21
where
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[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    23
  "alpha_res (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
    24
     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
    25
     (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
    26
     R (pi \<bullet> x) y"
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    28
fun
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    29
  alpha_lst
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    30
where
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[simp del]:
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    32
  "alpha_lst (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
    33
     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
    34
     (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
    35
     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
    36
     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
    37
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    38
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
    39
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    40
notation
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    41
  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
    42
  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
    43
  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
    44
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    45
section {* Mono *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    46
1557
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    47
lemma [mono]: 
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
    48
  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
    49
  and   "R1 \<le> R2 \<Longrightarrow> alpha_res bs R1 \<le> alpha_res bs R2"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    50
  and   "R1 \<le> R2 \<Longrightarrow> alpha_lst cs R1 \<le> alpha_lst cs R2"
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    51
  by (case_tac [!] bs, case_tac [!] cs) 
fee2389789ad moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents: 1544
diff changeset
    52
     (auto simp add: le_fun_def le_bool_def alphas)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    54
section {* Equivariance *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    55
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    56
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
    57
  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
    58
  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
    59
  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
    60
  unfolding alphas
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    61
  unfolding permute_eqvt[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    62
  unfolding set_eqvt[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    63
  unfolding permute_fun_app_eq[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    64
  unfolding Diff_eqvt[symmetric]
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    65
  by (auto simp add: permute_bool_def fresh_star_permute_iff)
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    66
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    67
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    68
section {* Equivalence *}
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    69
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    70
lemma alpha_refl:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    71
  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
    72
  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
    73
  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
    74
  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
    75
  using a 
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    76
  unfolding alphas
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    77
  unfolding fresh_star_def
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    78
  by (simp_all add: fresh_zero_perm)
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    79
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    80
lemma alpha_sym:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    81
  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
    82
  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
    83
  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
    84
  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
    85
  unfolding alphas fresh_star_def
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    86
  using a
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    87
  by (auto simp add:  fresh_minus_perm)
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    88
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    89
lemma alpha_trans:
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    90
  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
    91
  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
    92
  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
    93
  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
    94
  using a
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    95
  unfolding alphas fresh_star_def
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    96
  by (simp_all add: fresh_plus_perm)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    97
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
    98
lemma alpha_sym_eqvt:
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
    99
  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
   100
  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
   101
  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
   102
  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
   103
  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
   104
apply(auto intro!: alpha_sym)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   105
apply(drule_tac [!] a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   106
apply(rule_tac [!] p="p" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   107
apply(perm_simp add: permute_minus_cancel b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   108
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   109
apply(perm_simp add: permute_minus_cancel b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   110
apply(assumption)
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
done
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   114
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   115
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
   116
  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
   117
  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
   118
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   119
  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
   120
  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
   121
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   122
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   123
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   124
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   125
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   126
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   127
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   128
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   129
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   130
apply(drule_tac p="q" in permute_boolI)
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(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   133
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   134
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   135
lemma alpha_res_trans_eqvt:
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   136
  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
   137
  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
   138
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   139
  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
   140
  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
   141
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   142
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   143
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   144
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   145
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   146
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   147
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   148
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   149
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   150
apply(drule_tac p="q" in permute_boolI)
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(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   153
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   154
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   155
lemma alpha_lst_trans_eqvt:
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   156
  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
   157
  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
   158
  and     d: "q \<bullet> R = R"
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   159
  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
   160
  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
   161
apply(rule alpha_trans)
2313
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   162
defer
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   163
apply(rule a)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   164
apply(rule b)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   165
apply(drule c)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   166
apply(rule_tac p="q" in permute_boolE)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   167
apply(perm_simp add: permute_minus_cancel d)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   168
apply(assumption)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   169
apply(rotate_tac -1)
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   170
apply(drule_tac p="q" in permute_boolI)
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(simp add: permute_eqvt[symmetric])
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   173
done
25d2cdf7d7e4 transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents: 2311
diff changeset
   174
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   175
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
   176
2311
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   177
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   178
section {* General Abstractions *}
4da5c5c29009 work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents: 2302
diff changeset
   179
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
fun
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   181
  alpha_abs_set 
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   183
  [simp del]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   184
  "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
   185
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   186
fun
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   187
  alpha_abs_lst
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   188
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   189
  [simp del]:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   190
  "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
   191
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   192
fun
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   193
  alpha_abs_res
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   194
where
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   195
  [simp del]:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   196
  "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
   197
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
notation
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   199
  alpha_abs_set (infix "\<approx>abs'_set" 50) and
1666
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   200
  alpha_abs_lst (infix "\<approx>abs'_lst" 50) and
a99ae705b811 Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1665
diff changeset
   201
  alpha_abs_res (infix "\<approx>abs'_res" 50)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   202
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   203
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
   204
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   205
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   206
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
   207
  shows "(bs, x) \<approx>abs_set (bs, x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   208
  and   "(bs, x) \<approx>abs_res (bs, x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   209
  and   "(cs, x) \<approx>abs_lst (cs, x)" 
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   210
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   211
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   212
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   213
  by (rule_tac [!] x="0" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   214
     (simp_all add: fresh_zero_perm)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   215
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   216
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
   217
  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
   218
  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
   219
  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
   220
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   221
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   222
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   223
  by (erule_tac [!] exE, rule_tac [!] x="-p" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   224
     (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
   225
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   226
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
   227
  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
   228
  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
   229
  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
   230
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   231
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   232
  unfolding fresh_star_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   233
  apply(erule_tac [!] exE, erule_tac [!] exE)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   234
  apply(rule_tac [!] x="pa + p" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   235
  by (simp_all add: fresh_plus_perm)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   236
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   237
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
   238
  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
   239
  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
   240
  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
   241
  unfolding alphas_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   242
  unfolding alphas
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   243
  unfolding set_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   244
  unfolding supp_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   245
  unfolding Diff_eqvt[symmetric]
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   246
  apply(erule_tac [!] exE)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   247
  apply(rule_tac [!] x="p \<bullet> pa" in exI)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   248
  by (auto simp add: fresh_star_permute_iff permute_eqvt[symmetric])
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   249
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   250
quotient_type 
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   251
    'a abs_set = "(atom set \<times> 'a::pt)" / "alpha_abs_set"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   252
and 'b abs_res = "(atom set \<times> 'b::pt)" / "alpha_abs_res"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   253
and 'c abs_lst = "(atom list \<times> 'c::pt)" / "alpha_abs_lst"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   254
  apply(rule_tac [!] equivpI)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  unfolding reflp_def symp_def transp_def
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   256
  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
   257
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
quotient_definition
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   259
  Abs_set ("[_]set. _" [60, 60] 60)
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   260
where
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   261
  "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
   262
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   265
quotient_definition
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   266
  Abs_res ("[_]res. _" [60, 60] 60)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   267
where
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   268
  "Abs_res::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_res"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   269
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   270
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   271
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   272
quotient_definition
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   273
  Abs_lst ("[_]lst. _" [60, 60] 60)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   274
where
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   275
  "Abs_lst::atom list \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_lst"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   276
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   277
  "Pair::atom list \<Rightarrow> ('a::pt) \<Rightarrow> (atom list \<times> 'a)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   278
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
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
   280
  shows "(op= ===> op= ===> alpha_abs_set) Pair Pair"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   281
  and   "(op= ===> op= ===> alpha_abs_res) Pair Pair"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   282
  and   "(op= ===> op= ===> alpha_abs_lst) Pair Pair"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   283
  unfolding fun_rel_def
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   284
  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
   285
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
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
   287
  shows "(op= ===> alpha_abs_set ===> alpha_abs_set) permute permute"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   288
  and   "(op= ===> alpha_abs_res ===> alpha_abs_res) permute permute"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   289
  and   "(op= ===> alpha_abs_lst ===> alpha_abs_lst) permute permute"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   290
  unfolding fun_rel_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   291
  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
   292
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
   293
lemma abs_exhausts:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   294
  shows "(\<And>as (x::'a::pt). y1 = Abs_set as x \<Longrightarrow> P1) \<Longrightarrow> P1"
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
   295
  and   "(\<And>as (x::'a::pt). y2 = Abs_res as x \<Longrightarrow> P2) \<Longrightarrow> P2"
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
   296
  and   "(\<And>as (x::'a::pt). y3 = Abs_lst as x \<Longrightarrow> P3) \<Longrightarrow> P3"
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
   297
  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
   298
              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
   299
              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
   300
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   301
lemma abs_eq_iff:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   302
  shows "Abs_set bs x = Abs_set cs y \<longleftrightarrow> (bs, x) \<approx>abs_set (cs, y)"
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   303
  and   "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (bs, x) \<approx>abs_res (cs, y)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   304
  and   "Abs_lst ds x = Abs_lst hs y \<longleftrightarrow> (ds, x) \<approx>abs_lst (hs, y)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   305
  unfolding alphas_abs by (lifting alphas_abs)
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   306
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   307
instantiation abs_set :: (pt) pt
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
begin
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
quotient_definition
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   311
  "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
   312
is
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   315
lemma permute_Abs[simp]:
1558
a5ba76208983 started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents: 1557
diff changeset
   316
  fixes x::"'a::pt"  
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   317
  shows "(p \<bullet> (Abs_set as x)) = Abs_set (p \<bullet> as) (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   318
  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
   319
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
instance
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
  apply(default)
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
   322
  apply(case_tac [!] x rule: abs_exhausts(1))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   323
  apply(simp_all)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   324
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   325
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   326
end
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   327
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   328
instantiation abs_res :: (pt) pt
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   329
begin
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   330
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   331
quotient_definition
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   332
  "permute_abs_res::perm \<Rightarrow> ('a::pt abs_res) \<Rightarrow> 'a abs_res"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   333
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   334
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   335
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   336
lemma permute_Abs_res[simp]:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   337
  fixes x::"'a::pt"  
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   338
  shows "(p \<bullet> (Abs_res as x)) = Abs_res (p \<bullet> as) (p \<bullet> x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   339
  by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"])
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   340
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   341
instance
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   342
  apply(default)
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
   343
  apply(case_tac [!] x rule: abs_exhausts(2))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   344
  apply(simp_all)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   345
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   346
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   347
end
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   348
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   349
instantiation abs_lst :: (pt) pt
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   350
begin
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   351
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   352
quotient_definition
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   353
  "permute_abs_lst::perm \<Rightarrow> ('a::pt abs_lst) \<Rightarrow> 'a abs_lst"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   354
is
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   355
  "permute:: perm \<Rightarrow> (atom list \<times> 'a::pt) \<Rightarrow> (atom list \<times> 'a::pt)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   356
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   357
lemma permute_Abs_lst[simp]:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   358
  fixes x::"'a::pt"  
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   359
  shows "(p \<bullet> (Abs_lst as x)) = Abs_lst (p \<bullet> as) (p \<bullet> x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   360
  by (lifting permute_prod.simps[where 'a="atom list" and 'b="'a"])
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   361
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   362
instance
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   363
  apply(default)
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
   364
  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
   365
  apply(simp_all)
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
2460
16d32eddc17f added eqvt-attribute for permute_abs lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2447
diff changeset
   370
lemmas permute_abs[eqvt] = permute_Abs permute_Abs_res permute_Abs_lst
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   371
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   372
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   373
lemma abs_swap1:
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   374
  assumes a1: "a \<notin> (supp x) - bs"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   375
  and     a2: "b \<notin> (supp x) - bs"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   376
  shows "Abs_set bs x = Abs_set ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)"
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   377
  and   "Abs_res bs x = Abs_res ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   378
  unfolding abs_eq_iff
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   379
  unfolding alphas_abs
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   380
  unfolding alphas
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   381
  unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] 
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   382
  unfolding fresh_star_def fresh_def
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   383
  unfolding swap_set_not_in[OF a1 a2] 
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   384
  using a1 a2
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   385
  by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   386
     (auto simp add: supp_perm swap_atom)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   387
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   388
lemma abs_swap2:
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   389
  assumes a1: "a \<notin> (supp x) - (set bs)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   390
  and     a2: "b \<notin> (supp x) - (set bs)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   391
  shows "Abs_lst bs x = Abs_lst ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)"
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   392
  unfolding abs_eq_iff
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   393
  unfolding alphas_abs
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   394
  unfolding alphas
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   395
  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
   396
  unfolding fresh_star_def fresh_def
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   397
  unfolding swap_set_not_in[OF a1 a2]
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   398
  using a1 a2
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   399
  by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   400
     (auto simp add: supp_perm swap_atom)
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   401
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   402
lemma abs_supports:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   403
  shows "((supp x) - as) supports (Abs_set as x)"
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   404
  and   "((supp x) - as) supports (Abs_res as x)"
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   405
  and   "((supp x) - set bs) supports (Abs_lst bs x)"
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   406
  unfolding supports_def
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   407
  unfolding permute_abs
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1661
diff changeset
   408
  by (simp_all add: abs_swap1[symmetric] abs_swap2[symmetric])
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   409
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
   410
function
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   411
  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
   412
where
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   413
  "supp_set (Abs_set as x) = supp x - as"
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
   414
apply(case_tac x rule: abs_exhausts(1))
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
   415
apply(simp)
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
   416
apply(simp add: abs_eq_iff alphas_abs alphas)
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
   417
done
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   418
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   419
termination supp_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
   420
  by (auto intro!: local.termination)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
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
   422
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
   423
  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
   424
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
   425
  "supp_res (Abs_res as x) = supp x - as"
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
   426
apply(case_tac x rule: abs_exhausts(2))
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
   427
apply(simp)
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
   428
apply(simp add: abs_eq_iff alphas_abs alphas)
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
   429
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
   430
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
   431
termination supp_res 
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
   432
  by (auto intro!: local.termination)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
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
   434
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
   435
  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
   436
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
   437
  "supp_lst (Abs_lst cs x) = (supp x) - (set cs)"
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
   438
apply(case_tac x rule: abs_exhausts(3))
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
   439
apply(simp)
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
   440
apply(simp add: abs_eq_iff alphas_abs alphas)
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
   441
done
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
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
   443
termination supp_lst 
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
   444
  by (auto intro!: local.termination)
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
   445
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
   446
lemma [eqvt]:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   447
  shows "(p \<bullet> supp_set x) = supp_set (p \<bullet> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   448
  and   "(p \<bullet> supp_res y) = supp_res (p \<bullet> y)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   449
  and   "(p \<bullet> supp_lst z) = supp_lst (p \<bullet> z)"
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
   450
  apply(case_tac x rule: abs_exhausts(1))
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
   451
  apply(simp add: supp_eqvt Diff_eqvt)
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
   452
  apply(case_tac y rule: abs_exhausts(2))
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
   453
  apply(simp add: supp_eqvt Diff_eqvt)
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
   454
  apply(case_tac z rule: abs_exhausts(3))
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
   455
  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
   456
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   458
lemma aux_fresh:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   459
  shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_set (Abs bs x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   460
  and   "a \<sharp> Abs_res bs x \<Longrightarrow> a \<sharp> supp_res (Abs_res bs x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   461
  and   "a \<sharp> Abs_lst cs x \<Longrightarrow> a \<sharp> supp_lst (Abs_lst cs x)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   462
  by (rule_tac [!] fresh_fun_eqvt_app)
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   463
     (simp_all only: eqvts_raw)
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   464
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   465
lemma supp_abs_subset1:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   466
  assumes a: "finite (supp x)"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   467
  shows "(supp x) - as \<subseteq> supp (Abs_set as x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   468
  and   "(supp x) - as \<subseteq> supp (Abs_res as x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   469
  and   "(supp x) - (set bs) \<subseteq> supp (Abs_lst bs x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   470
  unfolding supp_conv_fresh
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   471
  by (auto dest!: aux_fresh)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   472
     (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
   473
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   474
lemma supp_abs_subset2:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   475
  assumes a: "finite (supp x)"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   476
  shows "supp (Abs_set as x) \<subseteq> (supp x) - as"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   477
  and   "supp (Abs_res as x) \<subseteq> (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   478
  and   "supp (Abs_lst bs x) \<subseteq> (supp x) - (set bs)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   479
  by (rule_tac [!] supp_is_subset)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   480
     (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
   481
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   482
lemma abs_finite_supp:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   483
  assumes a: "finite (supp x)"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   484
  shows "supp (Abs_set as x) = (supp x) - as"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   485
  and   "supp (Abs_res as x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   486
  and   "supp (Abs_lst bs x) = (supp x) - (set bs)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   487
  by (rule_tac [!] subset_antisym)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   488
     (simp_all add: supp_abs_subset1[OF a] supp_abs_subset2[OF a])
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   490
lemma supp_abs:
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  fixes x::"'a::fs"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   492
  shows "supp (Abs_set as x) = (supp x) - as"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   493
  and   "supp (Abs_res as x) = (supp x) - as"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   494
  and   "supp (Abs_lst bs x) = (supp x) - (set bs)"
1932
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   495
  by (rule_tac [!] abs_finite_supp)
2b0cc308fd6a tuned proofs
Christian Urban <urbanc@in.tum.de>
parents: 1924
diff changeset
   496
     (simp_all add: finite_supp)
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   498
instance abs_set :: (fs) fs
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
  apply(default)
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
   500
  apply(case_tac x rule: abs_exhausts(1))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   501
  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
   502
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   504
instance abs_res :: (fs) fs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   505
  apply(default)
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
   506
  apply(case_tac x rule: abs_exhausts(2))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   507
  apply(simp add: supp_abs finite_supp)
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   508
  done
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   509
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   510
instance abs_lst :: (fs) fs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   511
  apply(default)
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
   512
  apply(case_tac x rule: abs_exhausts(3))
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   513
  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
   514
  done
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   516
lemma abs_fresh_iff:
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   517
  fixes x::"'a::fs"
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   518
  shows "a \<sharp> Abs_set bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
1657
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   519
  and   "a \<sharp> Abs_res bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   520
  and   "a \<sharp> Abs_lst cs x \<longleftrightarrow> a \<in> (set cs) \<or> (a \<notin> (set cs) \<and> a \<sharp> x)"
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   521
  unfolding fresh_def
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   522
  unfolding supp_abs
Christian Urban <urbanc@in.tum.de>
parents: 1588
diff changeset
   523
  by auto
1460
0fd03936dedb merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents: 1451
diff changeset
   524
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   525
lemma Abs_eq_iff:
2469
4a6e78bd9de9 renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents: 2468
diff changeset
   526
  shows "Abs_set bs x = Abs_set cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op =) supp p (cs, y))"
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   527
  and   "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (cs, y))"
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   528
  and   "Abs_lst bsl x = Abs_lst csl y \<longleftrightarrow> (\<exists>p. (bsl, x) \<approx>lst (op =) supp p (csl, y))"
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   529
  by (lifting alphas_abs)
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   530
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   531
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2467
diff changeset
   532
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
   533
2385
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   534
fun
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   535
  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
   536
where
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   537
  "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
   538
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   539
definition 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   540
  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
   541
where
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   542
 "prod_alpha = prod_rel"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   543
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   544
lemma [quot_respect]:
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   545
  shows "((R1 ===> op =) ===> (R2 ===> op =) ===> prod_rel R1 R2 ===> op =) prod_fv prod_fv"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   546
  by auto
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   547
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   548
lemma [quot_preserve]:
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   549
  assumes q1: "Quotient R1 abs1 rep1"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   550
  and     q2: "Quotient R2 abs2 rep2"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   551
  shows "((abs1 ---> id) ---> (abs2 ---> id) ---> prod_fun rep1 rep2 ---> id) prod_fv = prod_fv"
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   552
  by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   553
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   554
lemma [mono]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   555
  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
   556
  unfolding prod_alpha_def
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   557
  by auto
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   558
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   559
lemma [eqvt]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   560
  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
   561
  unfolding prod_alpha_def
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   562
  unfolding prod_rel.simps
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   563
  by (perm_simp) (rule refl)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   564
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   565
lemma [eqvt]: 
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   566
  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
   567
  unfolding prod_fv.simps
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   568
  by (perm_simp) (rule refl)
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   569
fe25a3ffeb14 cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 2324
diff changeset
   570
1440
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
end
ffd5540ac2e9 added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572