Nominal/Ex/Beta.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 10 Nov 2011 01:15:19 +0000
changeset 3049 83744806b660
parent 3048 fc4b3e367c86
child 3054 da0fccee125c
permissions -rw-r--r--
proved supp for QVar; QApp still fails - probably stronger condistion is needed
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
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype lam =
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  Var "name"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| App "lam" "lam"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| 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
    13
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
section {* capture-avoiding substitution *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
nominal_primrec
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  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
    18
where
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  "(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
    20
| "(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
    21
| "\<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
    22
  unfolding eqvt_def subst_graph_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  apply (rule, perm_simp, rule)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  apply(rule TrueI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  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
    26
  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
    27
  apply(blast)+
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  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
    29
  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
    30
  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
    31
  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
    32
  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
    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
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
termination (eqvt)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  by lexicographic_order
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
lemma forget:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  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
    41
  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
    42
     (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
    43
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
lemma fresh_fact:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  fixes z::"name"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  assumes a: "atom z \<sharp> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
      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
    48
  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
    49
  using a b
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  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
    51
     (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
    52
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
lemma substitution_lemma:  
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  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
    55
  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
    56
using a 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
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
    58
   (auto simp add: fresh_fact forget)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
inductive
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  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
    62
where
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  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
    64
| equ_refl:  "t \<approx> t"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
| 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
    66
| 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
    67
| 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
    68
| 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
    69
| 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
    70
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
equivariance equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
nominal_inductive equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  avoids equ_beta: "x" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
       | equ_Lam: "x"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
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
    77
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  shows "(op = ===> equ) Var Var"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  and   "(equ ===> equ ===> equ) App App"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  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
    82
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
by (auto intro: equ.intros)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
lemma equ_subst1:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  assumes "t \<approx> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  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
    88
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
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
    90
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
apply(rule equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
apply(rule equ_beta)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
apply(simp add: fresh_fact)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
apply(subst (2) substitution_lemma)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
apply(simp add: fresh_at_base)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
apply(auto intro: equ_sym)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
apply(blast intro: equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
apply(simp add: equ_Lam)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
apply(simp add: equ_App1)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
apply(simp add: equ_App2)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
lemma equ_subst2:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  assumes "t \<approx> s"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  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
   109
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
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
   111
apply(simp add: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
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
   114
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
apply(metis equ_Lam)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  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
   120
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
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
   122
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  shows "(op = ===> equ ===> equ) permute permute"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
by (auto intro: eqvt)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
quotient_type qlam = lam / equ
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
apply(rule equivpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
apply(rule reflpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
apply(simp add: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
apply(rule sympI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
apply(simp add: equ_sym)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
apply(rule transpI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
apply(auto intro: equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
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
   139
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
   140
quotient_definition QLam ("QLam [_]._") 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  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
   142
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
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
   144
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
   145
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
instantiation qlam :: pt
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
begin
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
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
   150
  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
   151
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
instance
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
apply default
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
apply(simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
apply(rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
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
   165
 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
   166
 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
   167
 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
   168
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
lemma supports_abs_qlam:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  "(supp t) supports (abs_qlam t)"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
unfolding supports_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
unfolding fresh_def[symmetric]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
apply(perm_simp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
apply(simp add: swap_fresh_fresh)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
3049
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   178
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
   179
  "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
   180
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
   181
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   182
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
   183
  "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
   184
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
   185
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
   186
done
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   187
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   188
 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
   189
  "{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
   190
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
   191
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
   192
defer
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   193
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
   194
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
   195
done
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   196
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   197
 lemma supports_qapp:
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   198
  "supp (t, s) supports (QApp t s)"
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   199
apply(subgoal_tac "supp (t, s) supports (abs_qlam (App (rep_qlam t) (rep_qlam s)))")
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   200
apply(simp add: QApp_def)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   201
apply(subgoal_tac "supp (App (rep_qlam t) (rep_qlam s)) supports (abs_qlam (App (rep_qlam t) (rep_qlam s)))")
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   202
prefer 2
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   203
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
   204
apply(simp add: lam.supp)
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   205
oops
83744806b660 proved supp for QVar; QApp still fails - probably stronger condistion is needed
Christian Urban <urbanc@in.tum.de>
parents: 3048
diff changeset
   206
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
lemma "(p \<bullet> rep_qlam t) \<approx> rep_qlam (p \<bullet> t)"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
 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
   209
 apply (rule rep_abs_rsp[OF Quotient_qlam])
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
 apply (rule equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
 done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
section {* Supp *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
definition
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  "Supp t = \<Inter>{supp s | s. s \<approx> t}"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  shows "(equ ===> op=) Supp Supp"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
unfolding Supp_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
apply(rule allI)+
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
apply(rule impI)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
apply(rule_tac f="Inter" in arg_cong)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
apply (metis equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
by (metis equivp_def qlam_equivp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
quotient_definition "supp_qlam::qlam \<Rightarrow> atom set" 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  is "Supp::lam \<Rightarrow> atom set"
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
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  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
   235
  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
   236
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
   237
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
   238
thm QVar_def
3048
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
oops
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
 
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
section {* Size *}
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
definition
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  "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
   248
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
lemma [quot_respect]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  shows "(equ ===> op=) Size Size"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
unfolding fun_rel_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
unfolding Size_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
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
   255
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
apply (metis equ_trans)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
by (metis equivp_def qlam_equivp)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
instantiation qlam :: size
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
begin
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
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
   263
  is "Size::lam \<Rightarrow> nat"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
instance
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
apply default
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
thm lam.size
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  "size (QVar x) = 0"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
apply(simp add: Size_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
apply(rule Least_equality)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
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
   281
apply(auto intro: equ_refl)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
lemma
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  "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
   286
apply(descending)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
apply(simp add: Size_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
apply(rule_tac n="Suc (size t)" and m="size t" in Least_Suc2)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
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
   291
apply(auto intro: equ_refl)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
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
   294
apply(auto intro: equ_refl)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
defer
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
apply(simp add: Collect_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
apply(auto)[1]
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
apply(auto)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
term rep_qlam
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
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
   305
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
lemma [quot_preserve]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  assumes "Quotient equ Abs Rep"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  shows "(id ---> Rep ---> id) fresh = fresh"
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
using assms
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
unfolding Quotient_def
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
apply(simp add: map_fun_def)
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
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
   313
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
sorry
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
lemma [simp]:
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
  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
   318
  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
   319
  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
   320
apply(lifting subst.simps(1))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
apply(lifting subst.simps(2))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
apply(lifting subst.simps(3))
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
done
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
end
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
fc4b3e367c86 added initial test about alpha-beta-equated terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330