--- a/Quot/Examples/LamEx.thy Wed Jan 20 11:30:32 2010 +0100
+++ b/Quot/Examples/LamEx.thy Wed Jan 20 12:20:18 2010 +0100
@@ -483,16 +483,32 @@
sorry
function
- term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow> (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
+ term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
+ (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
+ ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
where
"term1_hom var app abs' (rVar x) = (var x)"
| "term1_hom var app abs' (rApp t u) =
app t u (term1_hom var app abs' t) (term1_hom var app abs' u)"
| "term1_hom var app abs' (rLam x u) =
abs' (\<lambda>y. [(x, y)] \<bullet> u) (\<lambda>y. term1_hom var app abs' ([(x, y)] \<bullet> u))"
-sorry
+apply(pat_completeness)
+apply(auto)
+done
-print_theorems
+lemma pi_size:
+ fixes pi::"name prm"
+ and t::"rlam"
+ shows "size (pi \<bullet> t) = size t"
+apply(induct t)
+apply(auto)
+done
+
+termination term1_hom
+ apply -
+ apply(relation "measure (\<lambda>(f1, f2, f3, t). size t)")
+apply(auto simp add: pi_size)
+done
lemma hom_res: "
\<forall>f_var. \<forall>f_app \<in> Respects(alpha ===> alpha ===> op =).