Quot/Nominal/Abs.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 17 Feb 2010 09:26:10 +0100
changeset 1164 fe0a31cf30a0
parent 1139 c4001cda9da3
child 1194 3d54fcc5f41a
permissions -rw-r--r--
Simplifying perm_eq
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
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
    77
(* introduced for examples *)
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    78
lemma alpha_gen_atom_sym:
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    79
  assumes a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
1026
278253330b6a Disambiguating the syntax.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1024
diff changeset
    80
  shows "\<exists>pi. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi ({atom b}, s) \<Longrightarrow>
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    81
        \<exists>pi. ({atom b}, s) \<approx>gen R f pi ({atom a}, t)"
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    82
  apply(erule exE)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    83
  apply(rule_tac x="- pi" in exI)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    84
  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
    85
  apply(erule conjE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    86
  apply(rule conjI)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    87
  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
    88
  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
    89
  apply simp
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    90
  apply(rule a)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    91
  apply assumption
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    92
  done
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    93
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    94
lemma alpha_gen_atom_trans:
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    95
  assumes a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
1026
278253330b6a Disambiguating the syntax.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1024
diff changeset
    96
  shows "\<lbrakk>\<exists>pi\<Colon>perm. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi ({atom aa}, ta);
1021
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    97
        \<exists>pi\<Colon>perm. ({atom aa}, ta) \<approx>gen R f pi ({atom ba}, sa)\<rbrakk>
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    98
    \<Longrightarrow> \<exists>pi\<Colon>perm. ({atom a}, t) \<approx>gen R f pi ({atom ba}, sa)"
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
    99
  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
   100
  apply(erule conjE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   101
  apply(erule exE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   102
  apply(erule conjE)+
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   103
  apply(rule_tac x="pia + pi" in exI)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   104
  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
   105
  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
   106
  apply(drule mp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   107
  apply(rotate_tac 4)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   108
  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
   109
  apply(simp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   110
  apply(rotate_tac 6)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   111
  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
   112
  apply(simp)
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   113
  done
bacf3584640e General alpha_gen_trans for one-variable abstraction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1015
diff changeset
   114
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   115
lemma alpha_gen_atom_eqvt:
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   116
  assumes a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
1026
278253330b6a Disambiguating the syntax.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1024
diff changeset
   117
  and     b: "\<exists>pia. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2)) f pia ({atom b}, s)"
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   118
  shows  "\<exists>pia. ({atom (pi \<bullet> a)}, pi \<bullet> t) \<approx>gen R f pia ({atom (pi \<bullet> b)}, pi \<bullet> s)"
1039
0d832c36b1bb fixed proofs in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1034
diff changeset
   119
  using b 
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(erule exE)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   122
  apply(rule_tac x="pi \<bullet> pia" in exI)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   123
  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
   124
  apply(erule conjE)+
1034
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1026
diff changeset
   125
  apply(rule conjI)
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   126
  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
1039
0d832c36b1bb fixed proofs in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1034
diff changeset
   127
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   128
  apply(rule conjI)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   129
  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
1039
0d832c36b1bb fixed proofs in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents: 1034
diff changeset
   130
  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
1024
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   131
  apply(subst permute_eqvt[symmetric])
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   132
  apply(simp)
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   133
  done
b3deb964ad26 Some equivariance machinery that comes useful in LF.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1021
diff changeset
   134
995
ee0619b5adff introduced a generic alpha (but not sure whether it is helpful)
Christian Urban <urbanc@in.tum.de>
parents: 989
diff changeset
   135
fun
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   136
  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
   137
where
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   138
  "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
   139
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   140
notation
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   141
  alpha_abs ("_ \<approx>abs _")
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   143
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
   144
  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
   145
  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
   146
  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
   147
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   148
  apply(rule_tac x="(a \<rightleftharpoons> b)" in exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   149
  apply(simp add: alpha_gen)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   150
  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   151
  apply(simp add: swap_set_not_in[OF a1 a2])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   152
  apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   153
  using a1 a2
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   154
  apply(simp add: fresh_star_def fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   155
  apply(blast)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   156
  apply(simp add: supp_swap)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   157
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   158
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   159
fun
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   160
  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
   161
where
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   162
  "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
   163
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   164
lemma supp_abs_fun_lemma:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   165
  assumes a: "x \<approx>abs y" 
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   166
  shows "supp_abs_fun x = supp_abs_fun y"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   167
  using a
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   168
  apply(induct rule: alpha_abs.induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   169
  apply(simp add: alpha_gen)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   170
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   171
  
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   172
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
   173
  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
   174
  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
   175
  apply(simp_all)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   176
  (* refl *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   177
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   178
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   179
  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
   180
  apply(simp)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   181
  (* symm *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   182
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   183
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   184
  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
   185
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   186
  apply(clarsimp)
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   187
  (* trans *)
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   188
  apply(clarify)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   189
  apply(rule exI)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   190
  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
   191
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   192
  apply(assumption)
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   193
  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
   194
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
quotient_definition
1139
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   197
  "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
   198
is
c4001cda9da3 renamed 'as' to 'is' everywhere.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1128
diff changeset
   199
  "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
   200
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
lemma [quot_respect]:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   202
  shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   203
  apply(clarsimp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   204
  apply(rule exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   205
  apply(rule alpha_gen_refl)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   206
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   207
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
lemma [quot_respect]:
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   210
  shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   211
  apply(clarsimp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   212
  apply(rule exI)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   213
  apply(rule alpha_gen_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   214
  apply(assumption)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   215
  apply(simp_all add: supp_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   216
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   217
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   218
lemma [quot_respect]:
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   219
  shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   220
  apply(simp add: supp_abs_fun_lemma)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   221
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   223
lemma abs_induct:
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   224
  "\<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
   225
  apply(lifting prod.induct[where 'a="atom set" and 'b="'a"])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   226
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
1089
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   228
(* TEST case *)
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   229
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
   230
thm abs_induct abs_induct2
66097fe4942a added a test case
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   231
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   232
instantiation abs :: (pt) pt
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
begin
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
quotient_definition
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   236
  "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
   237
is
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   238
  "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
   239
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
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
   241
  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
   242
  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
   243
  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
   244
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
instance
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  apply(default)
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   247
  apply(induct_tac [!] x rule: abs_induct)
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  apply(simp_all)
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  done
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
end
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   252
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   253
quotient_definition
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   254
  "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
   255
is
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   256
  "supp_abs_fun"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   257
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   258
lemma supp_Abs_fun_simp:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   259
  shows "supp_Abs_fun (Abs bs x) = (supp x) - bs"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   260
  by (lifting supp_abs_fun.simps(1))
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   261
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
   262
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
   263
  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
   264
  apply(induct_tac x rule: abs_induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   265
  apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   266
  done
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   267
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   268
lemma supp_Abs_fun_fresh:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   269
  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
   270
  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
   271
  apply(simp add: eqvts_raw)
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   272
  apply(simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   273
  done
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   274
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   275
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
   276
  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
   277
  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
   278
  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
   279
  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
   280
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   281
lemma Abs_supports:
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   282
  shows "((supp x) - as) supports (Abs as x)"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   283
  unfolding supports_def
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   284
  apply(clarify)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   285
  apply(simp (no_asm))
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   286
  apply(subst Abs_swap[symmetric])
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   287
  apply(simp_all)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   288
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   289
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   290
lemma supp_Abs_subset1:
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   291
  fixes x::"'a::fs"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   292
  shows "(supp x) - as \<subseteq> supp (Abs as x)"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   293
  apply(simp add: supp_conv_fresh)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   294
  apply(auto)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   295
  apply(drule_tac supp_Abs_fun_fresh)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   296
  apply(simp only: supp_Abs_fun_simp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   297
  apply(simp add: fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   298
  apply(simp add: supp_finite_atom_set finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   299
  done
1006
ef34da709a0b got rid of RAbst type - is now just pairs
Christian Urban <urbanc@in.tum.de>
parents: 1005
diff changeset
   300
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   301
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
   302
  fixes x::"'a::fs"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   303
  shows "supp (Abs as x) \<subseteq> (supp x) - as"
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   304
  apply(rule supp_is_subset)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   305
  apply(rule Abs_supports)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   306
  apply(simp add: finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   307
  done
986
98375dde48fc general abstraction operator and complete characterisation of its support and freshness
Christian Urban <urbanc@in.tum.de>
parents: 984
diff changeset
   308
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   309
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
   310
  fixes x::"'a::fs"
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   311
  shows "supp (Abs as x) = (supp x) - as"
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   312
  apply(rule_tac subset_antisym)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   313
  apply(rule supp_Abs_subset2)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   314
  apply(rule supp_Abs_subset1)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   315
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
1007
b4f956137114 renamed Abst/abst to Abs/abs
Christian Urban <urbanc@in.tum.de>
parents: 1006
diff changeset
   317
instance abs :: (fs) fs
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   318
  apply(default)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   319
  apply(induct_tac x rule: abs_induct)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   320
  apply(simp add: supp_Abs)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   321
  apply(simp add: finite_supp)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   322
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   324
lemma Abs_fresh_iff:
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  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
   326
  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
   327
  apply(simp add: fresh_def)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   328
  apply(simp add: supp_Abs)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   329
  apply(auto)
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   330
  done
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
1014
272ea46a1766 cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1007
diff changeset
   332
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
   333
  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
   334
  by (lifting alpha_abs.simps(1))
984
8e2dd0b29466 attempt of a general abstraction operator
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
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
   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
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
(* 
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
  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
   340
  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
   341
  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
   342
*)
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
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
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
   345
  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
   346
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
   347
  "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
   348
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
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
   350
  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
   351
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
   352
thm swap_set_not_in
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
   353
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
   354
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
   355
  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
   356
  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
   357
  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
   358
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
   359
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
   360
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
   361
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
   362
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
   363
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
   364
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
   365
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
   366
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
   367
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
   368
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
   369
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
   370
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
lemma alpha_abs_swap:
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
  assumes a1: "(supp x - bs) \<sharp>* p"
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
  and     a2: "(supp x - bs) \<sharp>* p"
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
  shows "(bs, x) \<approx>abs (p \<bullet> bs, 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
   375
  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
   376
  apply(rule_tac x="p" 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
   377
  apply(simp add: alpha_gen)
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
   378
  apply(simp add: supp_eqvt[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
   379
  apply(rule conjI)
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
   380
  apply(rule sym)
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
   381
  apply(rule 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
   382
  using a1 a2
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
   383
  apply(auto)[1]
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
   384
  oops
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
   385
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
   386
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
   387
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
   388
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
   389
  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
   390
  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
   391
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
   392
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
   393
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
   394
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
   395
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
   396
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
   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(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
   399
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
   400
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
   401
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
   402
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
   403
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
   404
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
   405
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
   406
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
   407
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
   408
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
   409
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
   410
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
   411
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
   412
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
   413
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
   414
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
   415
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
   416
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
   417
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
   418
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
   419
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
   420
thm supp_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
   421
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
   422
lemma perm_induct_test:
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
   423
  fixes P :: "perm => bool"
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
   424
  assumes zero: "P 0"
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
   425
  assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> 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
   426
  assumes plus: "\<And>p1 p2. \<lbrakk>supp (p1 + p2) = (supp p1 \<union> supp p2); P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
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
   427
  shows "P p"
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
   428
sorry
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
   429
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
   430
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
   431
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
   432
  "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = 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
   433
apply(induct p rule: perm_induct_test)
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
   434
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
   435
apply(rule swap_fresh_fresh)
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
   436
apply(case_tac "a \<in> supp 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
   437
apply(simp add: fresh_star_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
   438
apply(drule_tac x="a" in bspec)
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
   439
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
   440
apply(simp add: 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
   441
apply(simp add: supp_swap)
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
   442
apply(simp add: 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
   443
apply(case_tac "b \<in> supp 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
   444
apply(simp add: fresh_star_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
   445
apply(drule_tac x="b" in bspec)
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
   446
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
   447
apply(simp add: 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
   448
apply(simp add: supp_swap)
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
   449
apply(simp add: 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
   450
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
   451
apply(drule meta_mp)
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
   452
apply(simp add: 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
   453
apply(drule meta_mp)
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
   454
apply(simp add: 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
   455
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
   456
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
   457
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
   458
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
   459
  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
   460
  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
   461
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
   462
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
   463
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
   464
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
   465
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
   466
lemma
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
   467
  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of 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
   468
  shows "(a, x) \<approx>abs1 (b, 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
   469
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
   470
apply(case_tac "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
   471
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
   472
oops
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
   473
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
   474
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
   475
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
   476