author | Christian Urban <urbanc@in.tum.de> |
Mon, 17 Jan 2011 12:37:37 +0000 | |
changeset 2664 | a9a1ed3f5023 |
parent 2633 | d1d09177ec89 |
child 2678 | 494b859bfc16 |
permissions | -rw-r--r-- |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Height |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Nominal2" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
text {* |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
A small problem suggested by D. Wang. It shows how |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
the height of a lambda-terms behaves under substitution. |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
*} |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
atom_decl name |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
nominal_datatype lam = |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
Var "name" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
| App "lam" "lam" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
| Lam x::"name" l::"lam" bind x in l ("Lam [_]._" [100,100] 100) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
text {* Definition of the height-function on lambda-terms. *} |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
function |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
height :: "lam \<Rightarrow> int" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
where |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
"height (Var x) = 1" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
| "height (App t1 t2) = (max (height t1) (height t2)) + 1" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
| "height (Lam [a].t) = (height t) + 1" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
apply(rule_tac y="x" in lam.exhaust) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
apply(simp_all)[3] |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
sorry |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
termination |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
apply(relation "measure size") |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
apply(simp_all add: lam.size) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
done |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
text {* Definition of capture-avoiding substitution. *} |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
function |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
where |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
"(Var x)[y::=t'] = (if x=y then t' else (Var x))" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
| "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
| "\<lbrakk>atom x \<sharp> y; atom x \<sharp> t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
apply(case_tac x) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
apply(simp) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
apply(simp_all)[3] |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
apply(blast) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
apply(blast) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
apply(simp add: fresh_star_def fresh_Pair) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
apply(blast) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
apply(simp add: lam.distinct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
apply(simp add: lam.eq_iff) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
sorry |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
termination |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
apply(relation "measure (size o fst)") |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
apply(simp_all add: lam.size) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
done |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
text{* The next lemma is needed in the Var-case of the theorem below. *} |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
lemma height_ge_one: |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
shows "1 \<le> (height e)" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
apply(nominal_induct e rule: lam.strong_induct) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
apply(simp_all) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
done |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
text {* |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
Unlike the proplem suggested by Wang, however, the |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
theorem is here formulated entirely by using functions. |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
*} |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
theorem height_subst: |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
shows "height (e[x::=e']) \<le> ((height e) - 1) + (height e')" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
85 |
proof (nominal_induct e avoiding: x e' rule: lam.strong_induct) |
2633
d1d09177ec89
added proper case names for all induct and exhaust theorems
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
86 |
case (Var y) |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
have "1 \<le> height e'" by (rule height_ge_one) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
then show "height (Var y[x::=e']) \<le> height (Var y) - 1 + height e'" by simp |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
next |
2633
d1d09177ec89
added proper case names for all induct and exhaust theorems
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
90 |
case (Lam y e1) |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
have ih: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" by fact |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
moreover |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
have vc: "atom y \<sharp> x" "atom y \<sharp> e'" by fact+ (* usual variable convention *) |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
ultimately show "height ((Lam [y].e1)[x::=e']) \<le> height (Lam [y].e1) - 1 + height e'" by simp |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
95 |
next |
2633
d1d09177ec89
added proper case names for all induct and exhaust theorems
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
96 |
case (App e1 e2) |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
have ih1: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
and ih2: "height (e2[x::=e']) \<le> ((height e2) - 1) + (height e')" by fact+ |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by simp |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
qed |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
101 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
102 |
end |