author | Christian Urban <urbanc@in.tum.de> |
Tue, 05 Jul 2011 04:23:33 +0200 | |
changeset 2945 | 70bbd18ad194 |
parent 2678 | 494b859bfc16 |
child 3057 | d959bc9c800c |
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 |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
2 |
imports "Lambda" |
2632
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 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
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
|
12 |
shows "1 \<le> (height e)" |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
13 |
by (induct e rule: lam.induct) |
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
14 |
(simp_all) |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
theorem height_subst: |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
17 |
shows "height (e[x::=e']) \<le> height e - 1 + height e'" |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
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
|
19 |
case (Var y) |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
20 |
have "1 \<le> height e'" using height_ge_one by simp |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
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
|
22 |
next |
2633
d1d09177ec89
added proper case names for all induct and exhaust theorems
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
23 |
case (Lam y e1) |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
24 |
have ih: "height (e1[x::=e']) \<le> height e1 - 1 + height e'" by fact |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
moreover |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
have vc: "atom y \<sharp> x" "atom y \<sharp> e'" by fact+ (* usual variable convention *) |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
27 |
ultimately show "height ((Lam y e1)[x::=e']) \<le> height (Lam y e1) - 1 + height e'" by simp |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
next |
2633
d1d09177ec89
added proper case names for all induct and exhaust theorems
Christian Urban <urbanc@in.tum.de>
parents:
2632
diff
changeset
|
29 |
case (App e1 e2) |
2678
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
30 |
have ih1: "height (e1[x::=e']) \<le> (height e1) - 1 + height e'" |
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
31 |
and ih2: "height (e2[x::=e']) \<le> (height e2) - 1 + height e'" by fact+ |
494b859bfc16
defined height as a function that returns an integer
Christian Urban <urbanc@in.tum.de>
parents:
2633
diff
changeset
|
32 |
then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by simp |
2632
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
qed |
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
|
e8732350a29f
added small example for strong inductions; functions still need a sorry
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
end |