author | urbanc |
Thu, 01 Jun 2006 14:40:22 +0200 | |
changeset 19752 | 18e5aa65fb8b |
child 20396 | 4d0c33719348 |
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 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
3 |
(* Simple, but artificial, problem suggested by D. Wang *) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
4 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
5 |
theory Height |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
6 |
imports Lam_substs |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
7 |
(* |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
8 |
- inherit the type of alpha-equated lambda-terms, |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
9 |
the iteration combinator for this type and the |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
10 |
definition of capture-avoiding substitution |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
11 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
12 |
(the iteration combinator is not yet derived |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
13 |
automatically in the last stable version of |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
14 |
the nominal package) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
15 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
16 |
- capture-avoiding substitution is written |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
17 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
18 |
t[x::=t'] |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
19 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
20 |
and is defined such that |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
21 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
22 |
(Var y)[x::=t'] = (if x=y then t' else Var y) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
23 |
(App t1 t2)[x::=t'] = App (t1[x::=t']) (t2[x::=t']) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
24 |
y\<sharp>x \<and> y\<sharp>t2 \<Longrightarrow> (Lam [y].t)[x::=t'] = Lam [y].(t[x::=t']) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
25 |
*) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
26 |
begin |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
27 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
28 |
text {* definition of the height-function by cases *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
29 |
constdefs |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
30 |
height_Var :: "name \<Rightarrow> int" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
31 |
"height_Var \<equiv> \<lambda>(a::name). 1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
32 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
33 |
height_App :: "int \<Rightarrow> int \<Rightarrow> int" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
34 |
"height_App \<equiv> \<lambda>n1 n2. (max n1 n2)+1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
35 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
36 |
height_Lam :: "name \<Rightarrow> int \<Rightarrow> int" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
37 |
"height_Lam \<equiv> \<lambda>(a::name) n. n+1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
38 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
39 |
height :: "lam \<Rightarrow> int" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
40 |
"height \<equiv> itfun height_Var height_App height_Lam" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
41 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
42 |
text {* show that height is a function *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
43 |
lemma supp_height_Lam: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
44 |
shows "((supp height_Lam)::name set)={}" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
45 |
apply(simp add: height_Lam_def supp_def perm_fun_def perm_int_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
46 |
done |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
47 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
48 |
lemma fin_supp_height: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
49 |
shows "finite ((supp (height_Var,height_App,height_Lam))::name set)" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
50 |
by (finite_guess add: height_Var_def height_App_def height_Lam_def perm_int_def fs_name1) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
51 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
52 |
lemma FCB_height_Lam: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
53 |
shows "\<exists>(a::name). a\<sharp>height_Lam \<and> (\<forall>n. a\<sharp>height_Lam a n)" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
54 |
apply(simp add: height_Lam_def fresh_def supp_def perm_fun_def perm_int_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
55 |
done |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
56 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
57 |
text {* derive the characteristic equations for height from the iteration combinator *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
58 |
lemma height_Var: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
59 |
shows "height (Var c) = 1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
60 |
apply(simp add: height_def itfun_Var[OF fin_supp_height, OF FCB_height_Lam]) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
61 |
apply(simp add: height_Var_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
62 |
done |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
63 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
64 |
lemma height_App: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
65 |
shows "height (App t1 t2) = (max (height t1) (height t2))+1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
66 |
apply(simp add: height_def itfun_App[OF fin_supp_height, OF FCB_height_Lam]) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
67 |
apply(simp add: height_App_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
68 |
done |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
69 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
70 |
lemma height_Lam: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
71 |
shows "height (Lam [a].t) = (height t)+1" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
72 |
apply(simp add: height_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
73 |
apply(rule trans) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
74 |
apply(rule itfun_Lam[OF fin_supp_height, OF FCB_height_Lam]) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
75 |
apply(simp add: fresh_def supp_prod supp_height_Lam) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
76 |
apply(simp add: supp_def height_Var_def height_App_def perm_fun_def perm_int_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
77 |
apply(simp add: height_Lam_def) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
78 |
done |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
79 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
80 |
text {* add the characteristic equations of height to the simplifier *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
81 |
declare height_Var[simp] height_App[simp] height_Lam[simp] |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
82 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
83 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
84 |
text{* the next lemma is needed in the Var-case of the theorem *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
85 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
86 |
lemma height_ge_one: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
87 |
shows "1 \<le> (height e)" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
88 |
by (nominal_induct e rule: lam.induct) (simp | arith)+ |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
89 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
90 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
91 |
text {* unlike the proplem suggested by Wang, the theorem is formulated |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
92 |
here entirely by using functions *} |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
93 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
94 |
theorem height_subst: |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
95 |
shows "height (e[x::=e']) \<le> (((height e) - 1) + (height e'))" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
96 |
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
|
97 |
case (Var y) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
98 |
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
|
99 |
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
|
100 |
next |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
101 |
case (Lam y e1) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
102 |
hence ih: "height (e1[x::=e']) \<le> (((height e1) - 1) + (height e'))" by simp |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
103 |
moreover |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
104 |
have fresh: "y\<sharp>x" "y\<sharp>e'" by fact |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
105 |
ultimately show "height ((Lam [y].e1)[x::=e']) \<le> height (Lam [y].e1) - 1 + height e'" by simp |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
106 |
next |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
107 |
case (App e1 e2) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
108 |
hence ih1: "height (e1[x::=e']) \<le> (((height e1) - 1) + (height e'))" |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
109 |
and ih2: "height (e2[x::=e']) \<le> (((height e2) - 1) + (height e'))" by simp_all |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
110 |
then show "height ((App e1 e2)[x::=e']) \<le> height (App e1 e2) - 1 + height e'" by (simp, arith) |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
111 |
qed |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
112 |
|
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
113 |
end |
18e5aa65fb8b
added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff
changeset
|
114 |