Nominal/Abs.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 11 Mar 2010 19:41:11 +0100
changeset 1431 bc9ed52bcef5
parent 1423 d59f851926c5
child 1432 b41de1879dae
permissions -rw-r--r--
support of atoms at the end of Abs.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
989
af02b193a19a the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
Christian Urban <urbanc@in.tum.de>
parents: 988
diff changeset
     1
theory Abs
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1096
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "../Quotient" "../Quotient_Product"
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     5
(* the next three lemmas that should be in Nominal \<dots>\<dots>must be cleaned *)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     6
lemma ball_image: 
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     7
  shows "(\<forall>x \<in> p \<bullet> S. P x) = (\<forall>x \<in> S. P (p \<bullet> x))"
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     8
apply(auto)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     9
apply(drule_tac x="p \<bullet> x" in bspec)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    10
apply(simp add: mem_permute_iff)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    11
apply(simp)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    12
apply(drule_tac x="(- p) \<bullet> x" in bspec)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    13
apply(rule_tac p1="p" in mem_permute_iff[THEN iffD1])
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    14
apply(simp)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    15
apply(simp)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    16
done
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    17
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    18
lemma fresh_star_plus:
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    19
  fixes p q::perm
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    20
  shows "\<lbrakk>a \<sharp>* p;  a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)"
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    21
  unfolding fresh_star_def
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1063
diff changeset
    22
  by (simp add: fresh_plus_perm)
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    23
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    24
lemma fresh_star_permute_iff:
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    25
  shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x"
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    26
apply(simp add: fresh_star_def)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    27
apply(simp add: ball_image)
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    28
apply(simp add: fresh_permute_iff)
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    29
done
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
    30
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
    31
fun
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
    32
  alpha_gen 
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
where
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
    34
  alpha_gen[simp del]:
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    35
  "alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> f x - bs = f y - cs \<and> (f x - bs) \<sharp>* pi \<and> R (pi \<bullet> x) y"
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    36
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    37
notation
1026
278253330b6a Disambiguating the syntax.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1024
diff changeset
    38
  alpha_gen ("_ \<approx>gen _ _ _ _" [100, 100, 100, 100, 100] 100)
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    39
1005
9d5d9e7ff71b Monotonicity of ~~gen, needed for using it in inductive definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 995
diff changeset
    40
lemma [mono]: "R1 \<le> R2 \<Longrightarrow> alpha_gen x R1 \<le> alpha_gen x R2"
9d5d9e7ff71b Monotonicity of ~~gen, needed for using it in inductive definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 995
diff changeset
    41
  by (cases x) (auto simp add: le_fun_def le_bool_def alpha_gen.simps)
9d5d9e7ff71b Monotonicity of ~~gen, needed for using it in inductive definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 995
diff changeset
    42
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    43
lemma alpha_gen_refl:
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    44
  assumes a: "R x x"
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    45
  shows "(bs, x) \<approx>gen R f 0 (bs, x)"
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    46
  using a by (simp add: alpha_gen fresh_star_def fresh_zero_perm)
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    47
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    48
lemma alpha_gen_sym:
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    49
  assumes a: "(bs, x) \<approx>gen R f p (cs, y)"
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    50
  and     b: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x"
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    51
  shows "(cs, y) \<approx>gen R f (- p) (bs, x)"
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    52
  using a b by (simp add: alpha_gen fresh_star_def fresh_def supp_minus_perm)
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
    53
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    54
lemma alpha_gen_trans:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    55
  assumes a: "(bs, x) \<approx>gen R f p1 (cs, y)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    56
  and     b: "(cs, y) \<approx>gen R f p2 (ds, z)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    57
  and     c: "\<lbrakk>R (p1 \<bullet> x) y; R (p2 \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((p2 + p1) \<bullet> x) z"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    58
  shows "(bs, x) \<approx>gen R f (p2 + p1) (ds, z)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    59
  using a b c using supp_plus_perm
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    60
  apply(simp add: alpha_gen fresh_star_def fresh_def)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    61
  apply(blast)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    62
  done
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    63
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    64
lemma alpha_gen_eqvt:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    65
  assumes a: "(bs, x) \<approx>gen R f q (cs, y)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    66
  and     b: "R (q \<bullet> x) y \<Longrightarrow> R (p \<bullet> (q \<bullet> x)) (p \<bullet> y)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    67
  and     c: "p \<bullet> (f x) = f (p \<bullet> x)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    68
  and     d: "p \<bullet> (f y) = f (p \<bullet> y)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    69
  shows "(p \<bullet> bs, p \<bullet> x) \<approx>gen R f (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    70
  using a b
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    71
  apply(simp add: alpha_gen c[symmetric] d[symmetric] Diff_eqvt[symmetric])
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    72
  apply(simp add: permute_eqvt[symmetric])
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    73
  apply(simp add: fresh_star_permute_iff)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    74
  apply(clarsimp)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    75
  done
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
    76
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    77
lemma alpha_gen_compose_sym:
1301
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    78
  fixes pi
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    79
  assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)"
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    80
  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
1301
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    81
  shows "(ab, s) \<approx>gen R f (- pi) (aa, t)"
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    82
  using b apply -
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    83
  apply(simp add: alpha_gen.simps)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    84
  apply(erule conjE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    85
  apply(rule conjI)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    86
  apply(simp add: fresh_star_def fresh_minus_perm)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    87
  apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))")
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    88
  apply simp
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    89
  apply(rule a)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    90
  apply assumption
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    91
  done
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    92
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    93
lemma alpha_gen_compose_trans:
1301
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    94
  fixes pi pia
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    95
  assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)"
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    96
  and c: "(ab, ta) \<approx>gen R f pia (ac, sa)"
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    97
  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
1301
c145c380e195 Fix equivp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
    98
  shows "(aa, t) \<approx>gen R f (pia + pi) (ac, sa)"
1210
10a0e3578507 Generalize atom_trans and atom_sym.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1194
diff changeset
    99
  using b c apply -
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   100
  apply(simp add: alpha_gen.simps)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   101
  apply(erule conjE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   102
  apply(simp add: fresh_star_plus)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   103
  apply(drule_tac x="- pia \<bullet> sa" in spec)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   104
  apply(drule mp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   105
  apply(rotate_tac 4)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   106
  apply(drule_tac pi="- pia" in a)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   107
  apply(simp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   108
  apply(rotate_tac 6)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   109
  apply(drule_tac pi="pia" in a)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   110
  apply(simp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   111
  done
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   112
1255
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   113
lemma alpha_gen_compose_eqvt:
1300
22a084c9316b Fixed eqvt code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1259
diff changeset
   114
  fixes  pia
22a084c9316b Fixed eqvt code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1259
diff changeset
   115
  assumes b: "(g d, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2)) f pia (g e, s)"
1255
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   116
  and     c: "\<And>y. pi \<bullet> (g y) = g (pi \<bullet> y)"
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   117
  and     a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
1300
22a084c9316b Fixed eqvt code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1259
diff changeset
   118
  shows  "(g (pi \<bullet> d), pi \<bullet> t) \<approx>gen R f (pi \<bullet> pia) (g (pi \<bullet> e), pi \<bullet> s)"
1255
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   119
  using b
1039
0d832c36b1bb fixed proofs in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1034
diff changeset
   120
  apply -
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   121
  apply(simp add: alpha_gen.simps)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   122
  apply(erule conjE)+
1034
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1026
diff changeset
   123
  apply(rule conjI)
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   124
  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
1255
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   125
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric])
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   126
  apply(rule conjI)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   127
  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
1255
ab8ed83d0188 Simplified and finised eqvt proofs for t1 and t5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1210
diff changeset
   128
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric])
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   129
  apply(subst permute_eqvt[symmetric])
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   130
  apply(simp)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   131
  done
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   132
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
   133
fun
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   134
  alpha_abs 
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
   135
where
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   136
  "alpha_abs (bs, x) (cs, y) = (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   137
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   138
notation
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   139
  alpha_abs ("_ \<approx>abs _")
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   141
lemma alpha_abs_swap:
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   142
  assumes a1: "a \<notin> (supp x) - bs"
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   143
  and     a2: "b \<notin> (supp x) - bs"
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   144
  shows "(bs, x) \<approx>abs ((a \<rightleftharpoons> b) \<bullet> bs, (a \<rightleftharpoons> b) \<bullet> x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   145
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   146
  apply(rule_tac x="(a \<rightleftharpoons> b)" in exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   147
  apply(simp add: alpha_gen)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   148
  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   149
  apply(simp add: swap_set_not_in[OF a1 a2])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   150
  apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   151
  using a1 a2
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   152
  apply(simp add: fresh_star_def fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   153
  apply(blast)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   154
  apply(simp add: supp_swap)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   155
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   156
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   157
fun
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   158
  supp_abs_fun
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   159
where
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   160
  "supp_abs_fun (bs, x) = (supp x) - bs"
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   161
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   162
lemma supp_abs_fun_lemma:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   163
  assumes a: "x \<approx>abs y" 
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   164
  shows "supp_abs_fun x = supp_abs_fun y"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   165
  using a
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   166
  apply(induct rule: alpha_abs.induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   167
  apply(simp add: alpha_gen)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   168
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   169
  
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   170
quotient_type 'a abs = "(atom set \<times> 'a::pt)" / "alpha_abs"
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   171
  apply(rule equivpI)
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   172
  unfolding reflp_def symp_def transp_def
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   173
  apply(simp_all)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   174
  (* refl *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   175
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   176
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   177
  apply(rule alpha_gen_refl)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   178
  apply(simp)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   179
  (* symm *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   180
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   181
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   182
  apply(rule alpha_gen_sym)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   183
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   184
  apply(clarsimp)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   185
  (* trans *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   186
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   187
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   188
  apply(rule alpha_gen_trans)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   189
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   190
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   191
  apply(simp)
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   192
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
quotient_definition
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   195
  "Abs::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs"
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   196
is
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   197
  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
lemma [quot_respect]:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   200
  shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   201
  apply(clarsimp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   202
  apply(rule exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   203
  apply(rule alpha_gen_refl)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   204
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   205
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
lemma [quot_respect]:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   208
  shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   209
  apply(clarsimp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   210
  apply(rule exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   211
  apply(rule alpha_gen_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   212
  apply(assumption)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   213
  apply(simp_all add: supp_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   214
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   215
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   216
lemma [quot_respect]:
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   217
  shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   218
  apply(simp add: supp_abs_fun_lemma)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   219
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   221
lemma abs_induct:
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   222
  "\<lbrakk>\<And>as (x::'a::pt). P (Abs as x)\<rbrakk> \<Longrightarrow> P t"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   223
  apply(lifting prod.induct[where 'a="atom set" and 'b="'a"])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   224
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   226
(* TEST case *)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   227
lemmas abs_induct2 = prod.induct[where 'a="atom set" and 'b="'a::pt", quot_lifted]
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   228
thm abs_induct abs_induct2
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   229
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   230
instantiation abs :: (pt) pt
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
begin
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
quotient_definition
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   234
  "permute_abs::perm \<Rightarrow> ('a::pt abs) \<Rightarrow> 'a abs"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   235
is
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   236
  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
lemma permute_ABS [simp]:
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
   239
  fixes x::"'a::pt"  (* ??? has to be 'a \<dots> 'b does not work *)
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   240
  shows "(p \<bullet> (Abs as x)) = Abs (p \<bullet> as) (p \<bullet> x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   241
  by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"])
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
instance
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  apply(default)
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   245
  apply(induct_tac [!] x rule: abs_induct)
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  apply(simp_all)
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  done
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
end
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   250
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   251
quotient_definition
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   252
  "supp_Abs_fun :: ('a::pt) abs \<Rightarrow> atom \<Rightarrow> bool"
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   253
is
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   254
  "supp_abs_fun"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   255
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   256
lemma supp_Abs_fun_simp:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   257
  shows "supp_Abs_fun (Abs bs x) = (supp x) - bs"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   258
  by (lifting supp_abs_fun.simps(1))
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   259
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   260
lemma supp_Abs_fun_eqvt [eqvt]:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   261
  shows "(p \<bullet> supp_Abs_fun x) = supp_Abs_fun (p \<bullet> x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   262
  apply(induct_tac x rule: abs_induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   263
  apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   264
  done
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   265
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   266
lemma supp_Abs_fun_fresh:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   267
  shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_Abs_fun (Abs bs x)"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   268
  apply(rule fresh_fun_eqvt_app)
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   269
  apply(simp add: eqvts_raw)
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   270
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   271
  done
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   272
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   273
lemma Abs_swap:
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   274
  assumes a1: "a \<notin> (supp x) - bs"
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   275
  and     a2: "b \<notin> (supp x) - bs"
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   276
  shows "(Abs bs x) = (Abs ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x))"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   277
  using a1 a2 by (lifting alpha_abs_swap)
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   278
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   279
lemma Abs_supports:
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   280
  shows "((supp x) - as) supports (Abs as x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   281
  unfolding supports_def
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   282
  apply(clarify)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   283
  apply(simp (no_asm))
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   284
  apply(subst Abs_swap[symmetric])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   285
  apply(simp_all)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   286
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   287
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   288
lemma supp_Abs_subset1:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   289
  fixes x::"'a::fs"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   290
  shows "(supp x) - as \<subseteq> supp (Abs as x)"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   291
  apply(simp add: supp_conv_fresh)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   292
  apply(auto)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   293
  apply(drule_tac supp_Abs_fun_fresh)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   294
  apply(simp only: supp_Abs_fun_simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   295
  apply(simp add: fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   296
  apply(simp add: supp_finite_atom_set finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   297
  done
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   298
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   299
lemma supp_Abs_subset2:
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   300
  fixes x::"'a::fs"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   301
  shows "supp (Abs as x) \<subseteq> (supp x) - as"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   302
  apply(rule supp_is_subset)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   303
  apply(rule Abs_supports)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   304
  apply(simp add: finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   305
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   306
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   307
lemma supp_Abs:
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   308
  fixes x::"'a::fs"
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   309
  shows "supp (Abs as x) = (supp x) - as"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   310
  apply(rule_tac subset_antisym)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   311
  apply(rule supp_Abs_subset2)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   312
  apply(rule supp_Abs_subset1)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   313
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   315
instance abs :: (fs) fs
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   316
  apply(default)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   317
  apply(induct_tac x rule: abs_induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   318
  apply(simp add: supp_Abs)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   319
  apply(simp add: finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   320
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   322
lemma Abs_fresh_iff:
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  fixes x::"'a::fs"
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   324
  shows "a \<sharp> Abs bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   325
  apply(simp add: fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   326
  apply(simp add: supp_Abs)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   327
  apply(auto)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   328
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   330
lemma Abs_eq_iff:
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   331
  shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   332
  by (lifting alpha_abs.simps(1))
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
1015
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   334
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   335
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   336
(* 
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   337
  below is a construction site for showing that in the
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   338
  single-binder case, the old and new alpha equivalence 
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   339
  coincide
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   340
*)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   341
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   342
fun
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   343
  alpha1
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   344
where
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   345
  "alpha1 (a, x) (b, y) \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<rightleftharpoons> b) \<bullet> y \<and> a \<sharp> y)"
1015
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   346
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   347
notation 
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   348
  alpha1 ("_ \<approx>abs1 _")
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   349
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   350
fun
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   351
  alpha2
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   352
where
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   353
  "alpha2 (a, x) (b, y) \<longleftrightarrow> (\<exists>c. c \<sharp> (a,b,x,y) \<and> ((c \<rightleftharpoons> a) \<bullet> x) = ((c \<rightleftharpoons> b) \<bullet> y))"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   354
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   355
notation 
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   356
  alpha2 ("_ \<approx>abs2 _")
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   357
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   358
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   359
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   360
lemma qq:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   361
  fixes S::"atom set"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   362
  assumes a: "supp p \<inter> S = {}"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   363
  shows "p \<bullet> S = S"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   364
using a
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   365
apply(simp add: supp_perm permute_set_eq)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   366
apply(auto)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   367
apply(simp only: disjoint_iff_not_equal)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   368
apply(simp)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   369
apply (metis permute_atom_def_raw)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   370
apply(rule_tac x="(- p) \<bullet> x" in exI)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   371
apply(simp)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   372
apply(simp only: disjoint_iff_not_equal)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   373
apply(simp)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   374
apply(metis permute_minus_cancel)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   375
done
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   376
1015
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   377
lemma
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   378
  assumes a: "(a, x) \<approx>abs1 (b, y)" "sort_of a = sort_of b"
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   379
  shows "({a}, x) \<approx>abs ({b}, y)"
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   380
using a
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   381
apply(simp)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   382
apply(erule disjE)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   383
apply(simp)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   384
apply(rule exI)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   385
apply(rule alpha_gen_refl)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   386
apply(simp)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   387
apply(rule_tac x="(a \<rightleftharpoons> b)" in  exI)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   388
apply(simp add: alpha_gen)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   389
apply(simp add: fresh_def)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   390
apply(rule conjI)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   391
apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in  permute_eq_iff[THEN iffD1])
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   392
apply(rule trans)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   393
apply(simp add: Diff_eqvt supp_eqvt)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   394
apply(subst swap_set_not_in)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   395
back
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   396
apply(simp)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   397
apply(simp)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   398
apply(simp add: permute_set_eq)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   399
apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in fresh_star_permute_iff[THEN iffD1])
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   400
apply(simp add: permute_self)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   401
apply(simp add: Diff_eqvt supp_eqvt)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   402
apply(simp add: permute_set_eq)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   403
apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   404
apply(simp add: fresh_star_def fresh_def)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   405
apply(blast)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   406
apply(simp add: supp_swap)
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   407
done
683483199a5d added a single-binder alpha equivalence; showed one half of the equivalence proof between general and single binder case
Christian Urban <urbanc@in.tum.de>
parents: 1014
diff changeset
   408
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   409
lemma perm_zero:
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   410
  assumes a: "\<forall>x::atom. p \<bullet> x = x"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   411
  shows "p = 0"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   412
using a
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   413
by (simp add: expand_perm_eq)
1327
670701b19e8e added ACM style file for ICFP
Christian Urban <urbanc@in.tum.de>
parents: 1312
diff changeset
   414
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   415
fun
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   416
  add_perm 
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   417
where
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   418
  "add_perm [] = 0"
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   419
| "add_perm ((a, b) # xs) = (a \<rightleftharpoons> b) + add_perm xs"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   420
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   421
fun
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   422
  elem_perm
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   423
where
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   424
  "elem_perm [] = {}"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   425
| "elem_perm ((a, b) # xs) = {a, b} \<union> elem_perm xs"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   426
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   427
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   428
lemma add_perm_apend:
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   429
  shows "add_perm (xs @ ys) = add_perm xs + add_perm ys"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   430
apply(induct xs arbitrary: ys)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   431
apply(auto simp add: add_assoc)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   432
done
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   433
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   434
lemma perm_list_exists:
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   435
  fixes p::perm
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   436
  shows "\<exists>xs. p = add_perm xs \<and> supp xs \<subseteq> supp p"
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   437
apply(induct p taking: "\<lambda>p::perm. card (supp p)" rule: measure_induct)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   438
apply(rename_tac p)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   439
apply(case_tac "supp p = {}")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   440
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   441
apply(simp add: supp_perm)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   442
apply(drule perm_zero)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   443
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   444
apply(rule_tac x="[]" in exI)
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   445
apply(simp add: supp_Nil)
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   446
apply(subgoal_tac "\<exists>x. x \<in> supp p")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   447
defer
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   448
apply(auto)[1]
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   449
apply(erule exE)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   450
apply(drule_tac x="p + (((- p) \<bullet> x) \<rightleftharpoons> x)" in spec)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   451
apply(drule mp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   452
defer
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   453
apply(erule exE)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   454
apply(rule_tac x="xs @ [((- p) \<bullet> x, x)]" in exI)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   455
apply(simp add: add_perm_apend)
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   456
apply(erule conjE)
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   457
apply(drule sym)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   458
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   459
apply(simp add: expand_perm_eq)
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   460
apply(simp add: supp_append)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   461
apply(simp add: supp_perm supp_Cons supp_Nil supp_Pair supp_atom)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   462
apply(rule conjI)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   463
prefer 2
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   464
apply(auto)[1]
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   465
apply (metis left_minus minus_unique permute_atom_def_raw permute_minus_cancel(2))
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   466
defer
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   467
apply(rule psubset_card_mono)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   468
apply(simp add: finite_supp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   469
apply(rule psubsetI)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   470
defer
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   471
apply(subgoal_tac "x \<notin> supp (p + (- p \<bullet> x \<rightleftharpoons> x))")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   472
apply(blast)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   473
apply(simp add: supp_perm)
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   474
defer
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   475
apply(auto simp add: supp_perm)[1]
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   476
apply(case_tac "x = xa")
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   477
apply(simp)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   478
apply(case_tac "((- p) \<bullet> x) = xa")
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   479
apply(simp)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   480
apply(case_tac "sort_of xa = sort_of x")
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   481
apply(simp)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   482
apply(auto)[1]
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   483
apply(simp)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   484
apply(simp)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   485
apply(subgoal_tac "{a. p \<bullet> (- p \<bullet> x \<rightleftharpoons> x) \<bullet> a \<noteq> a} \<subseteq> {a. p \<bullet> a \<noteq> a}")
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   486
apply(blast)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   487
apply(auto simp add: supp_perm)[1]
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   488
apply(case_tac "x = xa")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   489
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   490
apply(case_tac "((- p) \<bullet> x) = xa")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   491
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   492
apply(case_tac "sort_of xa = sort_of x")
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   493
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   494
apply(auto)[1]
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   495
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   496
apply(simp)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   497
done
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   498
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   499
lemma tt0:
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   500
  fixes p::perm
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   501
  shows "(supp x) \<sharp>* p \<Longrightarrow> \<forall>a \<in> supp p. a \<sharp> x"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   502
apply(auto simp add: fresh_star_def supp_perm fresh_def)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   503
done
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   504
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   505
lemma uu0:
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   506
  shows "(\<forall>a \<in> elem_perm xs. a \<sharp> x) \<Longrightarrow> (add_perm xs \<bullet> x) = x"
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   507
apply(induct xs rule: add_perm.induct)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   508
apply(simp)
1351
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   509
apply(simp add: swap_fresh_fresh)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   510
done
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   511
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   512
lemma yy0:
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   513
  fixes xs::"(atom \<times> atom) list"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   514
  shows "supp xs = elem_perm xs"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   515
apply(induct xs rule: elem_perm.induct)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   516
apply(auto simp add: supp_Nil supp_Cons supp_Pair supp_atom)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   517
done
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   518
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   519
lemma tt1:
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   520
  shows "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   521
apply(drule tt0)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   522
apply(subgoal_tac "\<exists>xs. p = add_perm xs \<and> supp xs \<subseteq> supp p")
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   523
prefer 2
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   524
apply(rule perm_list_exists)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   525
apply(erule exE)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   526
apply(simp only: yy0)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   527
apply(rule uu0)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   528
apply(auto)
cffc5d78ab7f more proofs in Abs and work on Core Haskell
Christian Urban <urbanc@in.tum.de>
parents: 1350
diff changeset
   529
done
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   530
1350
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   531
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   532
lemma perm_induct_test:
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   533
  fixes P :: "perm => bool"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   534
  assumes fin: "finite (supp p)" 
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   535
  assumes zero: "P 0"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   536
  assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   537
  assumes plus: "\<And>p1 p2. \<lbrakk>supp p1 \<inter> supp p2 = {}; P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   538
  shows "P p"
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   539
using fin
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   540
apply(induct F\<equiv>"supp p" arbitrary: p rule: finite_induct)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   541
apply(simp add: supp_perm)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   542
apply(drule perm_zero)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   543
apply(simp add: zero)
5b31e49678fc added a lemma that permutations can be represented as sums of swapping
Christian Urban <urbanc@in.tum.de>
parents: 1327
diff changeset
   544
apply(rotate_tac 3)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   545
oops
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   546
lemma tt:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   547
  "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   548
oops
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   549
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   550
lemma yy:
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   551
  assumes "S1 - {x} = S2 - {x}" "x \<in> S1" "x \<in> S2"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   552
  shows "S1 = S2"
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   553
using assms
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   554
apply (metis insert_Diff_single insert_absorb)
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   555
done
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   556
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   557
lemma permute_boolI:
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   558
  fixes P::"bool"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   559
  shows "p \<bullet> P \<Longrightarrow> P"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   560
apply(simp add: permute_bool_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   561
done
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   562
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   563
lemma permute_boolE:
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   564
  fixes P::"bool"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   565
  shows "P \<Longrightarrow> p \<bullet> P"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   566
apply(simp add: permute_bool_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   567
done
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   568
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   569
lemma fresh_star_eqvt:
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   570
  shows "(p \<bullet> (as \<sharp>* x)) = (p \<bullet> as) \<sharp>* (p \<bullet> x)"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   571
apply(simp add: permute_bool_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   572
apply(auto simp add: fresh_star_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   573
apply(drule_tac x="(- p) \<bullet> xa" in bspec)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   574
apply(rule_tac p="p" in permute_boolI)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   575
apply(simp add: mem_eqvt)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   576
apply(rule_tac p="- p" in permute_boolI)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   577
apply(simp add: fresh_eqvt)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   578
apply(drule_tac x="p \<bullet> xa" in bspec)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   579
apply(rule_tac p="- p" in permute_boolI)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   580
apply(simp add: mem_eqvt)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   581
apply(rule_tac p="p" in permute_boolI)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   582
apply(simp add: fresh_eqvt)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   583
done
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   584
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   585
lemma kk:
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   586
  assumes a: "p \<bullet> x = y"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   587
  shows "\<forall>a \<in> supp x. (p \<bullet> a) \<in> supp y"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   588
using a
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   589
apply(auto)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   590
apply(rule_tac p="- p" in permute_boolI)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   591
apply(simp add: mem_eqvt supp_eqvt)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   592
done
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   593
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   594
lemma ww:
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   595
  assumes "a \<notin> supp x" "b \<in> supp x" "a \<noteq> b" "sort_of a = sort_of b"
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   596
  shows "((a \<rightleftharpoons> b) \<bullet> x) \<noteq> x"
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   597
apply(subgoal_tac "(supp x) supports x")
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   598
apply(simp add: supports_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   599
using assms
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   600
apply -
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   601
apply(drule_tac x="a" in spec)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   602
defer
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   603
apply(rule supp_supports)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   604
apply(auto)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   605
apply(rotate_tac 1)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   606
apply(drule_tac p="(a \<rightleftharpoons> b)" in permute_boolE)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   607
apply(simp add: mem_eqvt supp_eqvt)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   608
done
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   609
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   610
lemma zz:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   611
  assumes "p \<bullet> x \<noteq> p \<bullet> y"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   612
  shows "x \<noteq> y"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   613
using assms
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   614
apply(auto)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   615
done
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   616
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   617
lemma alpha_abs_sym:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   618
  assumes a: "({a}, x) \<approx>abs ({b}, y)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   619
  shows "({b}, y) \<approx>abs ({a}, x)"
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   620
using a
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   621
apply(simp)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   622
apply(erule exE)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   623
apply(rule_tac x="- p" in exI)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   624
apply(simp add: alpha_gen)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   625
apply(simp add: fresh_star_def fresh_minus_perm)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   626
apply (metis permute_minus_cancel(2))
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   627
done
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   628
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   629
lemma alpha_abs_trans:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   630
  assumes a: "({a1}, x1) \<approx>abs ({a2}, x2)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   631
  assumes b: "({a2}, x2) \<approx>abs ({a3}, x3)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   632
  shows "({a1}, x1) \<approx>abs ({a3}, x3)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   633
using a b
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   634
apply(simp)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   635
apply(erule exE)+
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   636
apply(rule_tac x="pa + p" in exI)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   637
apply(simp add: alpha_gen)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   638
apply(simp add: fresh_star_def fresh_plus_perm)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   639
done
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   640
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   641
lemma alpha_equal:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   642
  assumes a: "({a}, x) \<approx>abs ({a}, y)" 
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   643
  shows "(a, x) \<approx>abs1 (a, y)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   644
using a
1095
8441b4b2469d a few more attempts to show the equivalence between old and new way of defining alpha-equivalence
Christian Urban <urbanc@in.tum.de>
parents: 1089
diff changeset
   645
apply(simp)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   646
apply(erule exE)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   647
apply(simp add: alpha_gen)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   648
apply(erule conjE)+
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   649
apply(case_tac "a \<notin> supp x")
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   650
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   651
apply(subgoal_tac "supp x \<sharp>* p")
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   652
apply(drule tt1)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   653
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   654
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   655
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   656
apply(case_tac "a \<notin> supp y")
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   657
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   658
apply(drule tt1)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   659
apply(clarify)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   660
apply(simp (no_asm_use))
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   661
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   662
apply(simp)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   663
apply(drule yy)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   664
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   665
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   666
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   667
apply(case_tac "a \<sharp> p")
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   668
apply(subgoal_tac "supp y \<sharp>* p")
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   669
apply(drule tt1)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   670
apply(clarify)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   671
apply(simp (no_asm_use))
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   672
apply(metis)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   673
apply(auto simp add: fresh_star_def)[1]
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   674
apply(frule_tac kk)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   675
apply(drule_tac x="a" in bspec)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   676
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   677
apply(simp add: fresh_def)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   678
apply(simp add: supp_perm)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   679
apply(subgoal_tac "((p \<bullet> a) \<sharp> p)")
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   680
apply(simp add: fresh_def supp_perm)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   681
apply(simp add: fresh_star_def)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   682
done
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   683
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   684
lemma alpha_unequal:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   685
  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b" "a \<noteq> b"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   686
  shows "(a, x) \<approx>abs1 (b, y)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   687
using a
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   688
apply -
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   689
apply(subgoal_tac "a \<notin> supp x - {a}")
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   690
apply(subgoal_tac "b \<notin> supp x - {a}")
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   691
defer
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   692
apply(simp add: alpha_gen)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   693
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   694
apply(drule_tac alpha_abs_swap)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   695
apply(assumption)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   696
apply(simp only: insert_eqvt empty_eqvt swap_atom_simps)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   697
apply(drule alpha_abs_sym)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   698
apply(rotate_tac 4)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   699
apply(drule_tac alpha_abs_trans)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   700
apply(assumption)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   701
apply(drule alpha_equal)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   702
apply(simp)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   703
apply(rule_tac p="(a \<rightleftharpoons> b)" in permute_boolI)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   704
apply(simp add: fresh_eqvt)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   705
apply(simp add: fresh_def)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   706
done
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   707
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   708
lemma alpha_new_old:
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   709
  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b" 
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   710
  shows "(a, x) \<approx>abs1 (b, y)"
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   711
using a
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   712
apply(case_tac "a=b")
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   713
apply(simp only: alpha_equal)
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   714
apply(drule alpha_unequal)
1403
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   715
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   716
apply(simp)
4a10338c2535 almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
Christian Urban <urbanc@in.tum.de>
parents: 1351
diff changeset
   717
apply(simp)
1423
d59f851926c5 finally the proof that new and old alpha agree
Christian Urban <urbanc@in.tum.de>
parents: 1403
diff changeset
   718
done
1312
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   719
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   720
fun
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   721
  distinct_perms 
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   722
where
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   723
  "distinct_perms [] = True"
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   724
| "distinct_perms (p # ps) = (supp p \<inter> supp ps = {} \<and> distinct_perms ps)"
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   725
1431
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   726
(* support of concrete atom sets *)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   727
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   728
atom_decl name
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   729
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   730
lemma atom_eqvt_raw:
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   731
  fixes p::"perm"
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   732
  shows "(p \<bullet> atom) = atom"
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   733
by (simp add: expand_fun_eq permute_fun_def atom_eqvt)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   734
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   735
lemma atom_image_cong:
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   736
  shows "(atom ` X = atom ` Y) = (X = Y)"
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   737
apply(rule inj_image_eq_iff)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   738
apply(simp add: inj_on_def)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   739
done
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   740
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   741
lemma
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   742
  fixes as::"name set"
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   743
  shows "supp (atom ` as) = supp as"
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   744
apply(simp add: supp_def)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   745
apply(simp add: image_eqvt)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   746
apply(simp add: atom_eqvt_raw)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   747
apply(simp add: atom_image_cong)
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   748
done
bc9ed52bcef5 support of atoms at the end of Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1423
diff changeset
   749
1312
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   750
b0eae8c93314 added some more examples from Peter Sewell's bestiary
Christian Urban <urbanc@in.tum.de>
parents: 1307
diff changeset
   751
989
af02b193a19a the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
Christian Urban <urbanc@in.tum.de>
parents: 988
diff changeset
   752
end
988
a987b5acadc8 improved the proof slightly by defining alpha as a function and completely characterised the equality between two abstractions
Christian Urban <urbanc@in.tum.de>
parents: 986
diff changeset
   753