| author | krauss | 
| Tue, 13 May 2008 09:10:56 +0200 | |
| changeset 26877 | c3bb1f397811 | 
| parent 26648 | 25c07f3878b0 | 
| child 26966 | 071f40487734 | 
| permissions | -rw-r--r-- | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
1  | 
(* $Id$ *)  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
2  | 
|
| 25751 | 3  | 
theory Height  | 
4  | 
imports "../Nominal"  | 
|
5  | 
begin  | 
|
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
6  | 
|
| 25751 | 7  | 
text {*  
 | 
| 26262 | 8  | 
A small problem suggested by D. Wang. It shows how  | 
| 25751 | 9  | 
the height of a lambda-terms behaves under substitution.  | 
10  | 
*}  | 
|
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
11  | 
|
| 
20396
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
12  | 
atom_decl name  | 
| 
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
13  | 
|
| 
22829
 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 
urbanc 
parents: 
22418 
diff
changeset
 | 
14  | 
nominal_datatype lam =  | 
| 
 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 
urbanc 
parents: 
22418 
diff
changeset
 | 
15  | 
Var "name"  | 
| 
 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 
urbanc 
parents: 
22418 
diff
changeset
 | 
16  | 
| App "lam" "lam"  | 
| 
 
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
 
urbanc 
parents: 
22418 
diff
changeset
 | 
17  | 
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 | 
| 
20396
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
18  | 
|
| 25751 | 19  | 
text {* Definition of the height-function on lambda-terms. *} 
 | 
| 21555 | 20  | 
consts  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
21  | 
height :: "lam \<Rightarrow> int"  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
22  | 
|
| 21555 | 23  | 
nominal_primrec  | 
24  | 
"height (Var x) = 1"  | 
|
25  | 
"height (App t1 t2) = (max (height t1) (height t2)) + 1"  | 
|
26  | 
"height (Lam [a].t) = (height t) + 1"  | 
|
27  | 
apply(finite_guess add: perm_int_def)+  | 
|
28  | 
apply(rule TrueI)+  | 
|
29  | 
apply(simp add: fresh_int)  | 
|
30  | 
apply(fresh_guess add: perm_int_def)+  | 
|
31  | 
done  | 
|
| 21087 | 32  | 
|
| 25751 | 33  | 
text {* Definition of capture-avoiding substitution. *}
 | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
34  | 
|
| 21555 | 35  | 
consts  | 
36  | 
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
 | 
|
| 
20396
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
37  | 
|
| 21555 | 38  | 
nominal_primrec  | 
39  | 
"(Var x)[y::=t'] = (if x=y then t' else (Var x))"  | 
|
40  | 
"(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"  | 
|
41  | 
"\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"  | 
|
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21555 
diff
changeset
 | 
42  | 
apply(finite_guess)+  | 
| 21555 | 43  | 
apply(rule TrueI)+  | 
44  | 
apply(simp add: abs_fresh)  | 
|
| 
22418
 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 
urbanc 
parents: 
21555 
diff
changeset
 | 
45  | 
apply(fresh_guess)+  | 
| 
20396
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
46  | 
done  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
47  | 
|
| 26262 | 48  | 
text{* The next lemma is needed in the Var-case of the theorem below. *}
 | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
50  | 
lemma height_ge_one:  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
51  | 
shows "1 \<le> (height e)"  | 
| 25751 | 52  | 
by (nominal_induct e rule: lam.induct) (simp_all)  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
53  | 
|
| 25751 | 54  | 
text {* 
 | 
55  | 
Unlike the proplem suggested by Wang, however, the  | 
|
| 26262 | 56  | 
theorem is here formulated entirely by using functions.  | 
| 25751 | 57  | 
*}  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
59  | 
theorem height_subst:  | 
| 26262 | 60  | 
shows "height (e[x::=e']) \<le> ((height e) - 1) + (height e')"  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
61  | 
proof (nominal_induct e avoiding: x e' rule: lam.induct)  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
62  | 
case (Var y)  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
63  | 
have "1 \<le> height e'" by (rule height_ge_one)  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
64  | 
then show "height (Var y[x::=e']) \<le> height (Var y) - 1 + height e'" by simp  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
65  | 
next  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
66  | 
case (Lam y e1)  | 
| 26262 | 67  | 
hence ih: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')" by simp  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
68  | 
moreover  | 
| 23393 | 69  | 
have vc: "y\<sharp>x" "y\<sharp>e'" by fact+ (* usual variable convention *)  | 
| 21555 | 70  | 
ultimately show "height ((Lam [y].e1)[x::=e']) \<le> height (Lam [y].e1) - 1 + height e'" by simp  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
71  | 
next  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
72  | 
case (App e1 e2)  | 
| 26262 | 73  | 
hence ih1: "height (e1[x::=e']) \<le> ((height e1) - 1) + (height e')"  | 
74  | 
and ih2: "height (e2[x::=e']) \<le> ((height e2) - 1) + (height e')" by simp_all  | 
|
| 
20396
 
4d0c33719348
used the recursion combinator for the height and substitution function
 
urbanc 
parents: 
19752 
diff
changeset
 | 
75  | 
then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by simp  | 
| 
19752
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
76  | 
qed  | 
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
 
urbanc 
parents:  
diff
changeset
 | 
78  | 
end  |