1797
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1
theory Lambda
2454
+ − 2
imports "../Nominal2"
1594
+ − 3
begin
+ − 4
2784
+ − 5
+ − 6
(* inductive_cases do not simplify with simple equations *)
+ − 7
atom_decl var
+ − 8
nominal_datatype expr =
+ − 9
eCnst nat
+ − 10
| eVar nat
2779
+ − 11
2784
+ − 12
inductive
+ − 13
eval :: "expr \<Rightarrow> nat \<Rightarrow> bool"
+ − 14
where
+ − 15
evCnst: "eval (eCnst c) 0"
+ − 16
| evVar: "eval (eVar n) 2"
2779
+ − 17
2784
+ − 18
inductive_cases
+ − 19
evInv_Cnst: "eval (eCnst c) m"
+ − 20
+ − 21
thm evInv_Cnst[simplified expr.distinct expr.eq_iff]
+ − 22
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 23
1594
+ − 24
atom_decl name
+ − 25
2436
+ − 26
nominal_datatype lam =
1800
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 27
Var "name"
2436
+ − 28
| App "lam" "lam"
2683
+ − 29
| Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100)
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 30
2779
+ − 31
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 32
inductive
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 33
triv :: "lam \<Rightarrow> nat \<Rightarrow> bool"
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
where
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 35
Var: "triv (Var x) n"
2779
+ − 36
| App: "\<lbrakk>triv t1 n; triv t2 n\<rbrakk> \<Longrightarrow> triv (App t1 t2) n"
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 37
2777
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 38
lemma
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 39
"p \<bullet> (triv t x) = triv (p \<bullet> t) (p \<bullet> x)"
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 40
unfolding triv_def
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 41
apply(perm_simp)
2779
+ − 42
apply(rule refl)
2777
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 43
oops
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 44
(*apply(perm_simp)*)
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 45
2779
+ − 46
ML {*
+ − 47
Inductive.the_inductive @{context} "Lambda.triv"
+ − 48
*}
+ − 49
+ − 50
thm triv_def
+ − 51
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 52
equivariance triv
2768
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 53
nominal_inductive triv avoids Var: "{}::name set"
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 54
apply(auto simp add: fresh_star_def)
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 55
done
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 56
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 57
inductive
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 58
triv2 :: "lam \<Rightarrow> nat \<Rightarrow> bool"
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 59
where
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 60
Var1: "triv2 (Var x) 0"
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 61
| Var2: "triv2 (Var x) (n + n)"
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 62
| Var3: "triv2 (Var x) n"
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 63
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 64
equivariance triv2
639979b7fa6e
added permute_pure back into the nominal_inductive procedure; updated to Isabelle 17 April
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 65
nominal_inductive triv2 .
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 66
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 67
lemma Abs1_eq_fdest:
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 68
fixes x y :: "'a :: at_base"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 69
and S T :: "'b :: fs"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 70
assumes "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 71
and "x \<noteq> y \<Longrightarrow> atom y \<sharp> T \<Longrightarrow> atom x \<sharp> f x T"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 72
and "x \<noteq> y \<Longrightarrow> atom y \<sharp> T \<Longrightarrow> atom y \<sharp> f x T"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 73
and "x \<noteq> y \<Longrightarrow> atom y \<sharp> T \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> T = S \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> (f x T) = f y S"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 74
and "sort_of (atom x) = sort_of (atom y)"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 75
shows "f x T = f y S"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 76
using assms apply -
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 77
apply (subst (asm) Abs1_eq_iff')
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 78
apply simp_all
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 79
apply (elim conjE disjE)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 80
apply simp
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 81
apply(rule trans)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 82
apply (rule_tac p="(atom x \<rightleftharpoons> atom y)" in supp_perm_eq[symmetric])
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 83
apply(rule fresh_star_supp_conv)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 84
apply(simp add: supp_swap fresh_star_def)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 85
apply(simp add: swap_commute)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 86
done
2765
7ac5e5c86c7d
introduced framework for finetuning eqvt-rules; this solves problem with permute_pure called in nominal_inductive
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 87
2683
+ − 88
text {* height function *}
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 89
2666
+ − 90
nominal_primrec
2678
+ − 91
height :: "lam \<Rightarrow> int"
2666
+ − 92
where
2678
+ − 93
"height (Var x) = 1"
2683
+ − 94
| "height (App t1 t2) = max (height t1) (height t2) + 1"
+ − 95
| "height (Lam [x].t) = height t + 1"
2707
+ − 96
defer
2666
+ − 97
apply(rule_tac y="x" in lam.exhaust)
2683
+ − 98
apply(auto simp add: lam.distinct lam.eq_iff)
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 99
apply (erule Abs1_eq_fdest)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 100
apply(simp_all add: fresh_def pure_supp eqvt_at_def)
2707
+ − 101
apply(subgoal_tac "\<And>p x r. height_graph x r \<Longrightarrow> height_graph (p \<bullet> x) (p \<bullet> r)")
+ − 102
unfolding eqvt_def
+ − 103
apply(rule allI)
+ − 104
apply(simp add: permute_fun_def)
+ − 105
apply(rule ext)
+ − 106
apply(rule ext)
+ − 107
apply(simp add: permute_bool_def)
+ − 108
apply(rule iffI)
+ − 109
apply(drule_tac x="p" in meta_spec)
+ − 110
apply(drule_tac x="- p \<bullet> x" in meta_spec)
+ − 111
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
+ − 112
apply(simp)
+ − 113
apply(drule_tac x="-p" in meta_spec)
+ − 114
apply(drule_tac x="x" in meta_spec)
+ − 115
apply(drule_tac x="xa" in meta_spec)
+ − 116
apply(simp)
+ − 117
apply(erule height_graph.induct)
+ − 118
apply(perm_simp)
+ − 119
apply(rule height_graph.intros)
+ − 120
apply(perm_simp)
+ − 121
apply(rule height_graph.intros)
+ − 122
apply(assumption)
+ − 123
apply(assumption)
+ − 124
apply(perm_simp)
+ − 125
apply(rule height_graph.intros)
+ − 126
apply(assumption)
2666
+ − 127
done
+ − 128
+ − 129
termination
2683
+ − 130
by (relation "measure size") (simp_all add: lam.size)
2666
+ − 131
2707
+ − 132
thm height.simps
+ − 133
2683
+ − 134
+ − 135
text {* free name function - returns atom lists *}
2666
+ − 136
+ − 137
nominal_primrec
+ − 138
frees_lst :: "lam \<Rightarrow> atom list"
+ − 139
where
+ − 140
"frees_lst (Var x) = [atom x]"
2683
+ − 141
| "frees_lst (App t1 t2) = frees_lst t1 @ frees_lst t2"
+ − 142
| "frees_lst (Lam [x]. t) = removeAll (atom x) (frees_lst t)"
2707
+ − 143
defer
2666
+ − 144
apply(rule_tac y="x" in lam.exhaust)
+ − 145
apply(simp_all)[3]
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 146
apply(simp_all add: lam.distinct lam.eq_iff)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 147
apply (erule Abs1_eq_fdest)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 148
apply(simp add: supp_removeAll fresh_def)
2666
+ − 149
apply(drule supp_eqvt_at)
+ − 150
apply(simp add: finite_supp)
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 151
apply(auto simp add: fresh_def supp_removeAll eqvts eqvt_at_def)
2707
+ − 152
apply(subgoal_tac "\<And>p x r. frees_lst_graph x r \<Longrightarrow> frees_lst_graph (p \<bullet> x) (p \<bullet> r)")
+ − 153
unfolding eqvt_def
+ − 154
apply(rule allI)
+ − 155
apply(simp add: permute_fun_def)
+ − 156
apply(rule ext)
+ − 157
apply(rule ext)
+ − 158
apply(simp add: permute_bool_def)
+ − 159
apply(rule iffI)
+ − 160
apply(drule_tac x="p" in meta_spec)
+ − 161
apply(drule_tac x="- p \<bullet> x" in meta_spec)
+ − 162
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
+ − 163
apply(simp)
+ − 164
apply(drule_tac x="-p" in meta_spec)
+ − 165
apply(drule_tac x="x" in meta_spec)
+ − 166
apply(drule_tac x="xa" in meta_spec)
+ − 167
apply(simp)
+ − 168
apply(erule frees_lst_graph.induct)
+ − 169
apply(perm_simp)
+ − 170
apply(rule frees_lst_graph.intros)
+ − 171
apply(perm_simp)
+ − 172
apply(rule frees_lst_graph.intros)
+ − 173
apply(assumption)
+ − 174
apply(assumption)
+ − 175
apply(perm_simp)
+ − 176
apply(rule frees_lst_graph.intros)
+ − 177
apply(assumption)
2666
+ − 178
done
+ − 179
+ − 180
termination
+ − 181
apply(relation "measure size")
+ − 182
apply(simp_all add: lam.size)
+ − 183
done
+ − 184
2683
+ − 185
text {* a small test lemma *}
2666
+ − 186
lemma
2683
+ − 187
shows "supp t = set (frees_lst t)"
+ − 188
apply(induct t rule: frees_lst.induct)
2666
+ − 189
apply(simp_all add: lam.supp supp_at_base)
+ − 190
done
+ − 191
2683
+ − 192
text {* capture - avoiding substitution *}
+ − 193
2675
+ − 194
nominal_primrec
2683
+ − 195
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [90, 90, 90] 90)
2675
+ − 196
where
2683
+ − 197
"(Var x)[y ::= s] = (if x = y then s else (Var x))"
+ − 198
| "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])"
+ − 199
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])"
2707
+ − 200
defer
2683
+ − 201
apply(auto simp add: lam.distinct lam.eq_iff)
+ − 202
apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust)
+ − 203
apply(blast)+
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 204
apply(simp_all add: fresh_star_def fresh_Pair_elim)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 205
apply (erule Abs1_eq_fdest)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 206
apply(simp_all add: Abs_fresh_iff)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 207
apply(drule_tac a="atom (xa)" in fresh_eqvt_at)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 208
apply(simp_all add: finite_supp fresh_Pair)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 209
apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> sa = sa")
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 210
apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> ya = ya")
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 211
apply(simp add: eqvt_at_def)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 212
apply(rule perm_supp_eq,simp add: fresh_star_def fresh_Pair supp_swap)+
2707
+ − 213
apply(subgoal_tac "\<And>p x r. subst_graph x r \<Longrightarrow> subst_graph (p \<bullet> x) (p \<bullet> r)")
+ − 214
unfolding eqvt_def
+ − 215
apply(rule allI)
+ − 216
apply(simp add: permute_fun_def)
+ − 217
apply(rule ext)
+ − 218
apply(rule ext)
+ − 219
apply(simp add: permute_bool_def)
+ − 220
apply(rule iffI)
+ − 221
apply(drule_tac x="p" in meta_spec)
+ − 222
apply(drule_tac x="- p \<bullet> x" in meta_spec)
+ − 223
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
+ − 224
apply(simp)
+ − 225
apply(drule_tac x="-p" in meta_spec)
+ − 226
apply(drule_tac x="x" in meta_spec)
+ − 227
apply(drule_tac x="xa" in meta_spec)
+ − 228
apply(simp)
+ − 229
apply(erule subst_graph.induct)
+ − 230
apply(perm_simp)
+ − 231
apply(rule subst_graph.intros)
+ − 232
apply(perm_simp)
+ − 233
apply(rule subst_graph.intros)
+ − 234
apply(assumption)
+ − 235
apply(assumption)
+ − 236
apply(perm_simp)
+ − 237
apply(rule subst_graph.intros)
+ − 238
apply(simp add: fresh_Pair)
+ − 239
apply(assumption)
2675
+ − 240
done
+ − 241
2678
+ − 242
termination
2683
+ − 243
by (relation "measure (\<lambda>(t,_,_). size t)")
+ − 244
(simp_all add: lam.size)
+ − 245
+ − 246
lemma subst_eqvt[eqvt]:
+ − 247
shows "(p \<bullet> t[x ::= s]) = (p \<bullet> t)[(p \<bullet> x) ::= (p \<bullet> s)]"
+ − 248
by (induct t x s rule: subst.induct) (simp_all)
+ − 249
+ − 250
lemma forget:
+ − 251
shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
+ − 252
apply(nominal_induct t avoiding: x s rule: lam.strong_induct)
+ − 253
apply(auto simp add: lam.fresh fresh_at_base)
+ − 254
done
+ − 255
+ − 256
text {* same lemma but with subst.induction *}
+ − 257
lemma forget2:
+ − 258
shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
+ − 259
apply(induct t x s rule: subst.induct)
+ − 260
apply(auto simp add: lam.fresh fresh_at_base fresh_Pair)
+ − 261
done
+ − 262
+ − 263
lemma fresh_fact:
+ − 264
fixes z::"name"
+ − 265
assumes a: "atom z \<sharp> s"
+ − 266
and b: "z = y \<or> atom z \<sharp> t"
+ − 267
shows "atom z \<sharp> t[y ::= s]"
+ − 268
using a b
+ − 269
apply (nominal_induct t avoiding: z y s rule: lam.strong_induct)
+ − 270
apply (auto simp add: lam.fresh fresh_at_base)
+ − 271
done
+ − 272
+ − 273
lemma substitution_lemma:
+ − 274
assumes a: "x \<noteq> y" "atom x \<sharp> u"
+ − 275
shows "t[x ::= s][y ::= u] = t[y ::= u][x ::= s[y ::= u]]"
+ − 276
using a
+ − 277
by (nominal_induct t avoiding: x y s u rule: lam.strong_induct)
+ − 278
(auto simp add: fresh_fact forget)
+ − 279
+ − 280
lemma subst_rename:
+ − 281
assumes a: "atom y \<sharp> t"
+ − 282
shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet>t)[y ::= s]"
+ − 283
using a
+ − 284
apply (nominal_induct t avoiding: x y s rule: lam.strong_induct)
+ − 285
apply (auto simp add: lam.fresh fresh_at_base)
+ − 286
done
+ − 287
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 288
lemma height_ge_one:
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 289
shows "1 \<le> (height e)"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 290
by (induct e rule: lam.induct) (simp_all)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 291
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 292
theorem height_subst:
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 293
shows "height (e[x::=e']) \<le> ((height e) - 1) + (height e')"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 294
proof (nominal_induct e avoiding: x e' rule: lam.strong_induct)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 295
case (Var y)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 296
have "1 \<le> height e'" by (rule height_ge_one)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 297
then show "height (Var y[x::=e']) \<le> height (Var y) - 1 + height e'" by simp
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 298
next
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 299
case (Lam y e1)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 300
hence ih: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" by simp
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 301
moreover
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 302
have vc: "atom y\<sharp>x" "atom y\<sharp>e'" by fact+ (* usual variable convention *)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 303
ultimately show "height ((Lam [y]. e1)[x::=e']) \<le> height (Lam [y]. e1) - 1 + height e'" by simp
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 304
next
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 305
case (App e1 e2)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 306
hence ih1: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')"
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 307
and ih2: "height (e2[x::=e']) \<le> ((height e2) - 1) + (height e')" by simp_all
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 308
then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by simp
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 309
qed
2683
+ − 310
+ − 311
subsection {* single-step beta-reduction *}
+ − 312
+ − 313
inductive
+ − 314
beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>b _" [80,80] 80)
+ − 315
where
+ − 316
b1[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> App t1 s \<longrightarrow>b App t2 s"
+ − 317
| b2[intro]: "s1 \<longrightarrow>b s2 \<Longrightarrow> App t s1 \<longrightarrow>b App t s2"
+ − 318
| b3[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> Lam [x]. t1 \<longrightarrow>b Lam [x]. t2"
+ − 319
| b4[intro]: "atom x \<sharp> s \<Longrightarrow> App (Lam [x]. t) s \<longrightarrow>b t[x ::= s]"
+ − 320
+ − 321
equivariance beta
+ − 322
+ − 323
nominal_inductive beta
+ − 324
avoids b4: "x"
+ − 325
by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact)
+ − 326
+ − 327
text {* One-Reduction *}
+ − 328
+ − 329
inductive
+ − 330
One :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>1 _" [80,80] 80)
+ − 331
where
+ − 332
o1[intro]: "Var x \<longrightarrow>1 Var x"
+ − 333
| o2[intro]: "\<lbrakk>t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>1 App t2 s2"
+ − 334
| o3[intro]: "t1 \<longrightarrow>1 t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>1 Lam [x].t2"
+ − 335
| o4[intro]: "\<lbrakk>atom x \<sharp> (s1, s2); t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]"
+ − 336
+ − 337
equivariance One
+ − 338
+ − 339
nominal_inductive One
+ − 340
avoids o3: "x"
+ − 341
| o4: "x"
+ − 342
by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact)
+ − 343
+ − 344
lemma One_refl:
+ − 345
shows "t \<longrightarrow>1 t"
+ − 346
by (nominal_induct t rule: lam.strong_induct) (auto)
+ − 347
+ − 348
lemma One_subst:
+ − 349
assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2"
+ − 350
shows "t1[x ::= s1] \<longrightarrow>1 t2[x ::= s2]"
+ − 351
using a
+ − 352
apply(nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct)
+ − 353
apply(auto simp add: substitution_lemma fresh_at_base fresh_fact fresh_Pair)
+ − 354
done
+ − 355
+ − 356
lemma better_o4_intro:
+ − 357
assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2"
+ − 358
shows "App (Lam [x]. t1) s1 \<longrightarrow>1 t2[ x ::= s2]"
+ − 359
proof -
2685
+ − 360
obtain y::"name" where fs: "atom y \<sharp> (x, t1, s1, t2, s2)" by (rule obtain_fresh)
2683
+ − 361
have "App (Lam [x]. t1) s1 = App (Lam [y]. ((y \<leftrightarrow> x) \<bullet> t1)) s1" using fs
+ − 362
by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base)
+ − 363
also have "\<dots> \<longrightarrow>1 ((y \<leftrightarrow> x) \<bullet> t2)[y ::= s2]" using fs a by (auto simp add: One.eqvt)
+ − 364
also have "\<dots> = t2[x ::= s2]" using fs by (simp add: subst_rename[symmetric])
+ − 365
finally show "App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]" by simp
+ − 366
qed
+ − 367
+ − 368
+ − 369
+ − 370
section {* Locally Nameless Terms *}
2678
+ − 371
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 372
nominal_datatype ln =
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 373
LNBnd nat
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 374
| LNVar name
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 375
| LNApp ln ln
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 376
| LNLam ln
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 377
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 378
fun
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 379
lookup :: "name list \<Rightarrow> nat \<Rightarrow> name \<Rightarrow> ln"
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 380
where
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 381
"lookup [] n x = LNVar x"
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 382
| "lookup (y # ys) n x = (if x = y then LNBnd n else (lookup ys (n + 1) x))"
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 383
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 384
lemma [eqvt]:
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 385
shows "(p \<bullet> lookup xs n x) = lookup (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)"
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 386
by (induct xs arbitrary: n) (simp_all add: permute_pure)
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 387
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 388
nominal_primrec
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 389
trans :: "lam \<Rightarrow> name list \<Rightarrow> ln"
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 390
where
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 391
"trans (Var x) xs = lookup xs 0 x"
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 392
| "trans (App t1 t2) xs = LNApp (trans t1 xs) (trans t2 xs)"
2685
+ − 393
| "atom x \<sharp> xs \<Longrightarrow> trans (Lam [x]. t) xs = LNLam (trans t (x # xs))"
2707
+ − 394
defer
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 395
apply(case_tac x)
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 396
apply(simp)
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 397
apply(rule_tac y="a" and c="b" in lam.strong_exhaust)
2729
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 398
apply(simp_all add: fresh_star_def)[3]
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 399
apply(blast)
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 400
apply(blast)
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 401
apply(simp_all add: lam.distinct lam.eq_iff)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 402
apply(elim conjE)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 403
apply clarify
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 404
apply (erule Abs1_eq_fdest)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 405
apply (simp_all add: ln.fresh)
2675
+ − 406
prefer 2
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 407
apply(drule supp_eqvt_at)
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 408
apply (auto simp add: finite_supp supp_Pair fresh_def supp_Cons supp_at_base)[2]
2729
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 409
prefer 2
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 410
apply (subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> xsa = xsa")
2729
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 411
apply (simp add: eqvt_at_def)
2767
94f6f70e3067
New way of forward elimination of Abs1_eq and simplifications of the function obligation proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 412
apply (metis atom_name_def swap_fresh_fresh)
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 413
oops
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 414
2729
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 415
(* lemma helpr: "atom x \<sharp> ta \<Longrightarrow> Lam [xa]. ta = Lam [x]. ((xa \<leftrightarrow> x) \<bullet> ta)"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 416
apply (case_tac "x = xa")
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 417
apply simp
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 418
apply (simp add: lam.eq_iff Abs1_eq_iff flip_def[symmetric])
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 419
by (metis atom_eqvt flip_at_simps(2) fresh_permute_iff)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 420
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 421
lemma supp_lookup: "supp (lookup l n name) = {atom name} - supp l"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 422
apply (induct l arbitrary: n)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 423
apply (simp_all add: ln.supp supp_at_base supp_Nil supp_Cons pure_supp)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 424
done
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 425
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 426
lemma trans_eqvt[eqvt]: "p \<bullet> (trans t l) = trans (p \<bullet> t) (p \<bullet> l)"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 427
apply (induct t l rule: trans.induct)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 428
apply simp_all
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 429
apply (simp add: eqvts permute_pure)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 430
done
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 431
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 432
lemma diff_un: "a - (b \<union> c) = a - b - c"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 433
by blast
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 434
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 435
lemma supp_trans: "supp (trans t l) = supp t - supp l"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 436
apply (induct t arbitrary: l rule: lam.induct)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 437
apply (simp_all add: lam.supp supp_at_base supp_lookup ln.supp)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 438
apply blast
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 439
apply (rule_tac x="(lam, l)" and ?'a="name" in obtain_fresh)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 440
apply (simp add: fresh_Pair)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 441
apply clarify
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 442
apply (subgoal_tac "supp (Lambda.trans (Lam [a]. ((name \<leftrightarrow> a) \<bullet> lam)) l) =
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 443
supp lam - {atom name} - supp l")
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 444
using helpr
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 445
apply simp
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 446
apply (simp add: ln.supp)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 447
apply (subgoal_tac "supp ((name \<leftrightarrow> a) \<bullet> (Lambda.trans lam ((name \<leftrightarrow> a) \<bullet> (a # l)))) = supp lam - {atom name} - supp l")
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 448
apply (simp add: trans_eqvt)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 449
apply (simp add: supp_eqvt[symmetric])
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 450
apply (simp add: Diff_eqvt)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 451
apply (simp add: supp_eqvt supp_Cons union_eqvt)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 452
apply (simp add: diff_un)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 453
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1*})
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 454
apply (tactic {* Cong_Tac.cong_tac @{thm cong} 1*})
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 455
apply rule
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 456
prefer 2
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 457
apply rule
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 458
apply (simp add: supp_at_base)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 459
apply (subgoal_tac "(name \<leftrightarrow> a) \<bullet> (supp lam - {atom name}) = supp lam - {atom name}")
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 460
apply (simp add: eqvts)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 461
unfolding flip_def
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 462
apply (rule swap_fresh_fresh)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 463
apply (metis fresh_at_base fresh_def fresh_minus_atom_set lam.fsupp supp_at_base)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 464
by (metis fresh_def fresh_finite_atom_set fresh_minus_atom_set lam.fsupp)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 465
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 466
lemma "atom x \<sharp> trans_sumC (t, x # xsa)"
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 467
by (simp add: fresh_def meta_eq_to_obj_eq[OF trans_def, symmetric, unfolded fun_eq_iff] supp_trans supp_Cons supp_at_base)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 468
*)
337748e9b6b5
Reduce the definition of trans to FCB; test that FCB can be proved with simp rules.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 469
2667
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 470
nominal_datatype db =
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 471
DBVar nat
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 472
| DBApp db db
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 473
| DBLam db
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 474
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 475
abbreviation
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 476
mbind :: "'a option => ('a => 'b option) => 'b option" ("_ \<guillemotright>= _" [65,65] 65)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 477
where
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 478
"c \<guillemotright>= f \<equiv> case c of None => None | (Some v) => f v"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 479
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 480
lemma mbind_eqvt:
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 481
fixes c::"'a::pt option"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 482
shows "(p \<bullet> (c \<guillemotright>= f)) = ((p \<bullet> c) \<guillemotright>= (p \<bullet> f))"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 483
apply(cases c)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 484
apply(simp_all)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 485
apply(perm_simp)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 486
apply(rule refl)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 487
done
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 488
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 489
lemma mbind_eqvt_raw[eqvt_raw]:
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 490
shows "(p \<bullet> option_case) \<equiv> option_case"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 491
apply(rule eq_reflection)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 492
apply(rule ext)+
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 493
apply(case_tac xb)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 494
apply(simp_all)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 495
apply(rule_tac p="-p" in permute_boolE)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 496
apply(perm_simp add: permute_minus_cancel)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 497
apply(simp)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 498
apply(rule_tac p="-p" in permute_boolE)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 499
apply(perm_simp add: permute_minus_cancel)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 500
apply(simp)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 501
done
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 502
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 503
fun
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 504
index :: "atom list \<Rightarrow> nat \<Rightarrow> atom \<Rightarrow> nat option"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 505
where
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 506
"index [] n x = None"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 507
| "index (y # ys) n x = (if x = y then (Some n) else (index ys (n + 1) x))"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 508
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 509
lemma [eqvt]:
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 510
shows "(p \<bullet> index xs n x) = index (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 511
apply(induct xs arbitrary: n)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 512
apply(simp_all add: permute_pure)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 513
done
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 514
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 515
(*
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 516
nominal_primrec
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 517
trans :: "lam \<Rightarrow> atom list \<Rightarrow> db option"
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 518
where
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 519
"trans (Var x) xs = (index xs 0 (atom x) \<guillemotright>= (\<lambda>n. Some (DBVar n)))"
2707
+ − 520
| "trans (App t1 t2) xs = ((trans t1 xs) \<guillemotright>= (\<lambda>db1. (trans t2 xs) \<guillemotright>= (\<lambda>db2. Some (DBApp db1 db2))))"
+ − 521
| "trans (Lam [x].t) xs = (trans t (atom x # xs) \<guillemotright>= (\<lambda>db. Some (DBLam db)))"
2667
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 522
*)
e3f8673085b1
added a translation function from lambda-terms to deBruijn terms (equivariance fails at the moment)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 523
2675
+ − 524
2666
+ − 525
1594
+ − 526
end
+ − 527
+ − 528
+ − 529