Nominal/Ex/Beta.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 16 Dec 2011 16:01:12 +0900
changeset 3066 da60dc911055
parent 3064 ade2f8fcf8e8
child 3077 df1055004f52
permissions -rw-r--r--
Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Beta
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
  "../Nominal2"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
atom_decl name
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
     9
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
nominal_datatype lam =
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  Var "name"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| App "lam" "lam"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
| Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
section {* capture-avoiding substitution *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
nominal_primrec
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [90, 90, 90] 90)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
where
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  "(Var x)[y ::= s] = (if x = y then s else (Var x))"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
| "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
| "\<lbrakk>atom x \<sharp> y; atom x \<sharp> s\<rbrakk> \<Longrightarrow> (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  unfolding eqvt_def subst_graph_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  apply (rule, perm_simp, rule)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  apply(rule TrueI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  apply(auto simp add: lam.distinct lam.eq_iff)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  apply(blast)+
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  apply(simp_all add: fresh_star_def fresh_Pair_elim)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  apply(simp_all add: Abs_fresh_iff)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  apply(simp add: fresh_star_def fresh_Pair)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
termination (eqvt)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  by lexicographic_order
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
lemma forget:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  by (nominal_induct t avoiding: x s rule: lam.strong_induct)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
     (auto simp add: lam.fresh fresh_at_base)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
lemma fresh_fact:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  fixes z::"name"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  assumes a: "atom z \<sharp> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
      and b: "z = y \<or> atom z \<sharp> t"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  shows "atom z \<sharp> t[y ::= s]"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  using a b
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
     (auto simp add: lam.fresh fresh_at_base)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
lemma substitution_lemma:  
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  assumes a: "x \<noteq> y" "atom x \<sharp> u"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  shows "t[x ::= s][y ::= u] = t[y ::= u][x ::= s[y ::= u]]"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
using a 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
by (nominal_induct t avoiding: x y s u rule: lam.strong_induct)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
   (auto simp add: fresh_fact forget)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
inductive
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    62
  beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<rightarrow> _")
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    63
where
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    64
  beta_beta:  "atom x \<sharp> s \<Longrightarrow> App (Lam [x].t) s \<rightarrow> t[x ::= s]"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    65
| beta_Lam:   "t \<rightarrow> s \<Longrightarrow> Lam [x].t \<rightarrow> Lam [x].s"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    66
| beta_App1:  "t \<rightarrow> s \<Longrightarrow> App t u \<rightarrow> App s u"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    67
| beta_App2:  "t \<rightarrow> s \<Longrightarrow> App u t \<rightarrow> App u s"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    68
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    69
equivariance beta
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    70
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    71
nominal_inductive beta
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    72
  avoids beta_beta: "x" 
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    73
       | beta_Lam: "x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    74
by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    75
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    76
inductive
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  equ :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<approx> _")
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
where
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  equ_beta:  "atom x \<sharp> s \<Longrightarrow> App (Lam [x].t) s \<approx> t[x ::= s]"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
| equ_refl:  "t \<approx> t"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
| equ_sym:   "t \<approx> s \<Longrightarrow> s \<approx> t"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
| equ_trans: "\<lbrakk>t1 \<approx> t2; t2 \<approx> t3\<rbrakk> \<Longrightarrow> t1 \<approx> t3"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
| equ_Lam:   "t \<approx> s \<Longrightarrow> Lam [x].t \<approx> Lam [x].s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
| equ_App1:  "t \<approx> s \<Longrightarrow> App t u \<approx> App s u"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
| equ_App2:  "t \<approx> s \<Longrightarrow> App u t \<approx> App u s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
equivariance equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
nominal_inductive equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  avoids equ_beta: "x" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
       | equ_Lam: "x"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    94
inductive
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    95
  equ2 :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<approx>2 _")
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    96
where
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    97
  equ2_beta1:  "\<lbrakk>atom x \<sharp> (t2, s2); s1 \<approx>2 t1; t2 \<approx>2  s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) t2 \<approx>2 s1[x ::= s2]"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    98
| equ2_beta2:  "\<lbrakk>atom x \<sharp> (t2, s2); s1 \<approx>2 t1; t2 \<approx>2  s2\<rbrakk> \<Longrightarrow> s1[x ::= s2] \<approx>2 App (Lam [x].t1) t2"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
    99
| equ2_Var:   "Var x \<approx>2 Var x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   100
| equ2_Lam:   "t \<approx>2 s \<Longrightarrow> Lam [x].t \<approx>2 Lam [x].s"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   101
| equ2_App:   "\<lbrakk>t1 \<approx>2 s1; t2 \<approx>2 s2\<rbrakk> \<Longrightarrow> App t1 t2 \<approx>2 App s1 s2"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   102
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   103
equivariance equ2
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   104
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   105
nominal_inductive equ2
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   106
  avoids equ2_beta1: "x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   107
       | equ2_beta2: "x" 
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   108
       | equ2_Lam: "x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   109
by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   110
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   111
lemma equ2_refl:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   112
  fixes t::"lam"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   113
  shows "t \<approx>2 t"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   114
apply(induct t rule: lam.induct)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   115
apply(auto intro: equ2.intros)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   116
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   117
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   118
lemma equ2_symm:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   119
  fixes t s::"lam"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   120
  assumes "t \<approx>2 s"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   121
  shows "s \<approx>2 t"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   122
using assms
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   123
apply(induct)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   124
apply(auto intro: equ2.intros)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   125
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   126
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   127
lemma equ2_trans:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   128
  fixes t1 t2 t3::"lam"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   129
  assumes "t1 \<approx>2 t2" "t2 \<approx>2 t3"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   130
  shows "t1 \<approx>2 t3"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   131
using assms
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   132
apply(nominal_induct t1 t2 avoiding: t3 rule: equ2.strong_induct)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   133
defer
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   134
defer
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   135
apply(erule equ2.cases)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   136
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   137
apply(rule equ2_beta2)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   138
apply(simp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   139
apply(simp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   140
apply(simp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   141
apply(rule equ2_refl)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   142
defer
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   143
apply(rotate_tac 4)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   144
apply(erule equ2.cases)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   145
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   146
oops
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   147
3066
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   148
lemma "a \<noteq> x \<Longrightarrow> x \<noteq> y \<Longrightarrow> y \<noteq> a \<Longrightarrow> y \<noteq> b \<Longrightarrow> (App (App (Lam[x]. Lam[y]. (App (Var y) (Var x))) (Var a)) (Var b) \<approx> App (Var b) (Var a))"
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   149
  apply (rule equ_trans)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   150
  apply (rule equ_App1)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   151
  apply (rule equ_beta)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   152
  apply (simp add: lam.fresh fresh_at_base)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   153
  apply (subst subst.simps)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   154
  apply (simp add: lam.fresh fresh_at_base)+
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   155
  apply (rule equ_trans)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   156
  apply (rule equ_beta)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   157
  apply (simp add: lam.fresh fresh_at_base)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   158
  apply (simp add: equ_refl)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   159
  done
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   160
3066
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   161
lemma "\<not> (Var b \<approx>2 Var a) \<Longrightarrow> \<not> (App (App (Lam[x]. Lam[y]. (App (Var y) (Var x))) (Var a)) (Var b) \<approx>2 App (Var b) (Var a))"
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   162
  apply rule
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   163
  apply (erule equ2.cases)
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   164
  apply auto
da60dc911055 Beta: equ and equ2 are not the same relations, equ2 seems not to be beta-eta equality.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 3064
diff changeset
   165
  done
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   166
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  shows "(op = ===> equ) Var Var"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  and   "(equ ===> equ ===> equ) App App"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  and   "(op = ===> equ ===> equ) Beta.Lam Beta.Lam"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
by (auto intro: equ.intros)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
lemma equ_subst1:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  assumes "t \<approx> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  shows "t[x ::= u] \<approx> s[x ::= u]"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
apply(nominal_induct avoiding: x u rule: equ.strong_induct)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
apply(rule equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
apply(rule equ_beta)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
apply(simp add: fresh_fact)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
apply(subst (2) substitution_lemma)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
apply(simp add: fresh_at_base)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
apply(auto intro: equ_sym)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
apply(blast intro: equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
apply(simp add: equ_Lam)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
apply(simp add: equ_App1)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
apply(simp add: equ_App2)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
lemma equ_subst2:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  assumes "t \<approx> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  shows "u[x ::= t] \<approx> u[x ::= s]"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
apply(nominal_induct u avoiding: x t s rule: lam.strong_induct)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
apply(simp add: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
apply(smt equ_App1 equ_App2 equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
apply(metis equ_Lam)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  shows "(equ ===> op = ===> equ ===> equ) subst subst"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
by (metis equ_subst1 equ_subst2 equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  shows "(op = ===> equ ===> equ) permute permute"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
unfolding fun_rel_def
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   215
by (auto intro: eqvts)
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
quotient_type qlam = lam / equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
apply(rule equivpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
apply(rule reflpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
apply(simp add: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
apply(rule sympI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
apply(simp add: equ_sym)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
apply(rule transpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
apply(auto intro: equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
quotient_definition "QVar::name \<Rightarrow> qlam" is Var
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
quotient_definition "QApp::qlam \<Rightarrow> qlam \<Rightarrow> qlam" is App
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
quotient_definition QLam ("QLam [_]._") 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  where "QLam::name \<Rightarrow> qlam \<Rightarrow> qlam" is Beta.Lam
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
lemmas qlam_strong_induct = lam.strong_induct[quot_lifted]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
lemmas qlam_induct = lam.induct[quot_lifted]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
instantiation qlam :: pt
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
begin
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
quotient_definition "permute_qlam::perm \<Rightarrow> qlam \<Rightarrow> qlam" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  is "permute::perm \<Rightarrow> lam \<Rightarrow> lam"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
instance
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
apply default
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   253
lemma qlam_perm[simp]:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   254
  shows "p \<bullet> (QVar x) = QVar (p \<bullet> x)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   255
  and "p \<bullet> (QApp t s) = QApp (p \<bullet> t) (p \<bullet> s)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   256
  and "p \<bullet> (QLam [x].t) = QLam [p \<bullet> x].(p \<bullet> t)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   257
apply(descending)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   258
apply(simp add: equ_refl)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   259
apply(descending)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   260
apply(simp add: equ_refl)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   261
apply(descending)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   262
apply(simp add: equ_refl)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   263
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   264
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   265
lemma qlam_supports:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   266
  shows "{atom x} supports (QVar x)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   267
  and   "supp (t, s) supports (QApp t s)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   268
  and   "supp (x, t) supports (QLam [x].t)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   269
unfolding supports_def fresh_def[symmetric]
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   270
by (auto simp add: swap_fresh_fresh)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   271
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   272
lemma qlam_fs:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   273
  fixes t::"qlam"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   274
  shows "finite (supp t)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   275
apply(induct t rule: qlam_induct)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   276
apply(rule supports_finite)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   277
apply(rule qlam_supports)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   278
apply(simp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   279
apply(rule supports_finite)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   280
apply(rule qlam_supports)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   281
apply(simp add: supp_Pair)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   282
apply(rule supports_finite)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   283
apply(rule qlam_supports)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   284
apply(simp add: supp_Pair finite_supp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   285
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   286
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   287
instantiation qlam :: fs
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   288
begin
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   289
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   290
instance
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   291
apply(default)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   292
apply(rule qlam_fs)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   293
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   294
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   295
end
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   296
3054
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   297
quotient_definition subst_qlam ("_[_ ::q= _]")
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   298
  where "subst_qlam::qlam \<Rightarrow> name \<Rightarrow> qlam \<Rightarrow> qlam" is subst  
3054
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   299
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   300
lemma
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   301
  "(QVar x)[y ::q= s] = (if x = y then s else (QVar x))"
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   302
apply(descending)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   303
apply(simp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   304
apply(auto intro: equ_refl)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   305
done
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   306
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   307
lemma
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   308
  "(QApp t1 t2)[y ::q= s] = QApp (t1[y ::q= s]) (t2[y ::q= s])"
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   309
apply(descending)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   310
apply(simp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   311
apply(rule equ_refl)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   312
done
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   313
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   314
definition
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   315
  "Supp t = \<Inter>{supp s | s. s \<approx> t}"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   316
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   317
lemma [quot_respect]:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   318
  shows "(equ ===> op=) Supp Supp"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   319
unfolding fun_rel_def
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   320
unfolding Supp_def
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   321
apply(rule allI)+
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   322
apply(rule impI)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   323
apply(rule_tac f="Inter" in arg_cong)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   324
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   325
apply (metis equ_trans)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   326
by (metis equivp_def qlam_equivp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   327
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   328
quotient_definition "supp_qlam::qlam \<Rightarrow> atom set" 
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   329
  is "Supp::lam \<Rightarrow> atom set"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   330
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   331
lemma Supp_supp:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   332
  "Supp t \<subseteq> supp t"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   333
unfolding Supp_def
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   334
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   335
apply(drule_tac x="supp t" in spec)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   336
apply(auto simp add: equ_refl)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   337
done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   338
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   339
lemma Supp_Lam:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   340
  "atom a \<notin> Supp (Lam [a].t)"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   341
proof -
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   342
  have "atom a \<notin> supp (Lam [a].t)" by (simp add: lam.supp)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   343
  then show ?thesis using Supp_supp
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   344
    by blast
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   345
qed
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   346
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   347
lemmas s = Supp_Lam[quot_lifted]
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   348
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   349
definition
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   350
  ssupp :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ ssupp _")
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   351
where
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   352
  "A ssupp x \<equiv> \<forall>p. (\<forall>a \<in> A. (p \<bullet> a) = a) \<longleftrightarrow> (p \<bullet> x) = x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   353
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   354
lemma ssupp_supports:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   355
  "A ssupp x \<Longrightarrow> A supports x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   356
unfolding ssupp_def supports_def
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   357
apply(rule allI)+
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   358
apply(drule_tac x="(a \<rightleftharpoons> b)" in spec)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   359
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   360
by (metis swap_atom_simps(3))
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   361
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   362
lemma ssupp_supp:
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   363
  assumes a: "finite A"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   364
  and     b: "A ssupp x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   365
  shows "supp x = A"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   366
proof -
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   367
  { assume 0: "A - supp x \<noteq> {}"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   368
    then obtain a where 1: "a \<in> A - supp x" by auto
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   369
    obtain a' where *: "a' \<in>  UNIV - A" and **: "sort_of a' = sort_of a"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   370
      by (rule obtain_atom[OF a]) (auto)  
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   371
    have "(a \<rightleftharpoons> a') \<bullet> a = a'" using 1 ** * by (auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   372
    then have w0: "(a \<rightleftharpoons> a') \<bullet> x \<noteq> x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   373
      using b unfolding ssupp_def 
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   374
      apply -
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   375
      apply(drule_tac x="(a \<rightleftharpoons> a')" in spec)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   376
      apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   377
      using 1 *
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   378
      apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   379
      done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   380
    have w1: "a \<sharp> x" unfolding fresh_def using 1 by auto 
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   381
    have w2: "a' \<sharp> x" using *
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   382
      apply(rule_tac supports_fresh)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   383
      apply(rule ssupp_supports)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   384
      apply(rule b)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   385
      apply(rule a)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   386
      apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   387
      done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   388
    have "False" using w0 w1 w2 by (simp add: swap_fresh_fresh)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   389
    then have "supp x = A" by simp }
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   390
  moreover
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   391
  { assume "A - supp x = {}"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   392
    moreover
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   393
    have "supp x \<subseteq> A"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   394
      apply(rule supp_is_subset)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   395
      apply(rule ssupp_supports)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   396
      apply(rule b)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   397
      apply(rule a)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   398
      done
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   399
    ultimately have "supp x = A"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   400
      by blast
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   401
  }
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   402
  ultimately show "supp x = A" by blast
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   403
qed
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   404
      
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   405
3054
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   406
lemma
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   407
  "(supp x) ssupp x"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   408
unfolding ssupp_def
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   409
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   410
apply(rule supp_perm_eq)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   411
apply(simp add: fresh_star_def)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   412
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   413
apply(simp add: fresh_perm)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   414
apply(simp add: fresh_perm[symmetric])
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   415
(*Tzevelekos 2008, section 2.1.2 property 2.5*)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   416
oops
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   417
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   418
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   419
function
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   420
  qsubst :: "qlam \<Rightarrow> name \<Rightarrow> qlam \<Rightarrow> qlam"  ("_ [_ ::qq= _]" [90, 90, 90] 90)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   421
where
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   422
  "(QVar x)[y ::qq= s] = (if x = y then s else (QVar x))"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   423
| "(QApp t1 t2)[y ::qq= s] = QApp (t1[y ::qq= s]) (t2[y ::qq= s])"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   424
| "\<lbrakk>atom x \<sharp> y; atom x \<sharp> s\<rbrakk> \<Longrightarrow> (QLam [x]. t)[y ::qq= s] = QLam [x].(t[y ::qq= s])"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   425
apply(simp_all add:)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   426
oops
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   427
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   428
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   429
lemma
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   430
  assumes a: "Lam [x].t \<approx> s"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   431
  shows "\<exists>x' t'. s = Lam [x']. t' \<and> t \<approx> t'"
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   432
using a
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   433
apply(induct s rule:lam.induct)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   434
apply(erule equ.cases)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   435
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   436
apply(erule equ.cases)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   437
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   438
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   439
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   440
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   441
lemma
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   442
  "atom x \<sharp> y \<Longrightarrow> atom x \<notin> supp_qlam s \<Longrightarrow> (QLam [x].t)[y ::q= s] = QLam [x].(t[y ::q= s])"
3054
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   443
apply(descending)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   444
oops
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   445
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   446
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
lemma [eqvt]: "(p \<bullet> abs_qlam t) = abs_qlam (p \<bullet> t)"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
 apply (subst fun_cong[OF fun_cong[OF meta_eq_to_obj_eq[OF permute_qlam_def]], of p, simplified])
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
 apply (subst Quotient_rel[OF Quotient_qlam, simplified equivp_reflp[OF qlam_equivp], simplified])
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
 by (metis Quotient_qlam equ_refl eqvt rep_abs_rsp_left)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
lemma supports_abs_qlam:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
  "(supp t) supports (abs_qlam t)"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
unfolding supports_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
unfolding fresh_def[symmetric]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
apply(perm_simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
apply(simp add: swap_fresh_fresh)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
3049
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   461
lemma rep_qlam_inverse:
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   462
  "abs_qlam (rep_qlam t) = t"
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   463
by (metis Quotient_abs_rep Quotient_qlam)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   464
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   465
lemma supports_rep_qlam:
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   466
  "supp (rep_qlam t) supports t"
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   467
apply(subst (2) rep_qlam_inverse[symmetric])
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   468
apply(rule supports_abs_qlam)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   469
done
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   470
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   471
 lemma supports_qvar:
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   472
  "{atom x} supports (QVar x)"
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   473
apply(subgoal_tac "supp (Var x) supports (abs_qlam (Var x))")
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   474
apply(simp add: lam.supp supp_at_base)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   475
defer
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   476
apply(rule supports_abs_qlam)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   477
apply(simp add: QVar_def)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   478
done
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   479
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
section {* Supp *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
3054
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   482
lemma s:
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   483
  assumes "s \<approx> t"
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   484
  shows "a \<sharp> (abs_qlam s) \<longleftrightarrow> a \<sharp> (abs_qlam t)"
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   485
using assms
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   486
by (metis Quotient_qlam Quotient_rel_abs)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   487
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   488
lemma ss:
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   489
  "Supp t supports (abs_qlam t)"
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   490
apply(simp only: supports_def)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   491
apply(rule allI)+
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   492
apply(rule impI)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   493
apply(rule swap_fresh_fresh)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   494
apply(drule conjunct1)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   495
apply(auto)[1]
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   496
apply(simp add: Supp_def)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   497
apply(auto)[1]
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   498
apply(simp add: s[symmetric])
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   499
apply(rule supports_fresh)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   500
apply(rule supports_abs_qlam)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   501
apply(simp add: finite_supp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   502
apply(simp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   503
apply(drule conjunct2)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   504
apply(auto)[1]
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   505
apply(simp add: Supp_def)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   506
apply(auto)[1]
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   507
apply(simp add: s[symmetric])
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   508
apply(rule supports_fresh)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   509
apply(rule supports_abs_qlam)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   510
apply(simp add: finite_supp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   511
apply(simp)
da0fccee125c a few more experiments with alpha-beta
Christian Urban <urbanc@in.tum.de>
parents: 3049
diff changeset
   512
done
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
  fixes t::"qlam"
3049
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   516
  shows "(supp x) supports (QVar x)"
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   517
apply(rule_tac x="QVar x" in Abs_qlam_cases)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   518
apply(auto)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   519
thm QVar_def
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
oops
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
section {* Size *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
definition
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  "Size t = Least {size s | s. s \<approx> t}" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  shows "(equ ===> op=) Size Size"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
unfolding Size_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
apply(rule_tac f="Least" in arg_cong)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
apply (metis equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
by (metis equivp_def qlam_equivp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
instantiation qlam :: size
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
begin
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
quotient_definition "size_qlam::qlam \<Rightarrow> nat" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
  is "Size::lam \<Rightarrow> nat"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
instance
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
apply default
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
thm lam.size
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
  "size (QVar x) = 0"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
apply(simp add: Size_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
apply(rule Least_equality)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
apply(rule_tac x="Var x" in exI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
apply(auto intro: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  "size (QLam [x].t) = Suc (size t)"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
apply(simp add: Size_def)
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   568
thm Least_Suc
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   569
apply(rule trans)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   570
apply(rule_tac n="Suc (size t)" in Least_Suc)
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
apply(rule_tac x="Lam [x].t" in exI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
apply(auto intro: equ_refl)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
apply(simp add: Collect_def)
3064
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   575
apply(auto)
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   576
defer
ade2f8fcf8e8 a bit more on alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents: 3054
diff changeset
   577
apply(simp add: Collect_def)
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
apply(rule_tac x="t" in exI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
apply(auto intro: equ_refl)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
defer
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
apply(auto)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
term rep_qlam
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
lemmas qlam_size_def = Size_def[quot_lifted]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
lemma [quot_preserve]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
  assumes "Quotient equ Abs Rep"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
  shows "(id ---> Rep ---> id) fresh = fresh"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
unfolding Quotient_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
apply(simp add: map_fun_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
apply(simp add: comp_def fun_eq_iff)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
sorry
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
lemma [simp]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
  shows "(QVar x)[y :::= s] = (if x = y then s else (QVar x))"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  and "(QApp t1 t2)[y :::= s] = QApp (t1[y :::= s]) (t2[y :::= s])"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  and "\<lbrakk>atom x \<sharp> y; atom x \<sharp> s\<rbrakk> \<Longrightarrow> (QLam [x]. t)[y :::= s] = QLam [x].(t[y :::= s])"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
apply(lifting subst.simps(1))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
apply(lifting subst.simps(2))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
apply(lifting subst.simps(3))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615