src/HOL/Nominal/Examples/Height.thy
author haftmann
Mon, 02 Oct 2006 23:00:45 +0200
changeset 20831 4981b56f8cde
parent 20398 1db76dd407bb
child 21053 7d0962594902
permissions -rw-r--r--
restructured contents
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
     6
imports Nominal
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
     7
begin
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
     8
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
     9
atom_decl name
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    10
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    11
nominal_datatype lam = Var "name"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    12
                     | App "lam" "lam"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    13
                     | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    14
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    15
thm lam.recs
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    16
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    17
types 'a f1_ty  = "name\<Rightarrow>('a::pt_name)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    18
      'a f2_ty  = "lam\<Rightarrow>lam\<Rightarrow>'a\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    19
      'a f3_ty  = "name\<Rightarrow>lam\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    20
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    21
text {* definition of the height-function by "structural recursion" ;o) *} 
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    22
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    23
constdefs 
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    24
  height_Var :: "name \<Rightarrow> int"
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    25
  "height_Var \<equiv> \<lambda>_. 1"
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    26
  
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    27
  height_App :: "lam\<Rightarrow>lam\<Rightarrow>int\<Rightarrow>int\<Rightarrow>int"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    28
  "height_App \<equiv> \<lambda>_ _ n1 n2. (max n1 n2)+1"
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    29
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    30
  height_Lam :: "name\<Rightarrow>lam\<Rightarrow>int\<Rightarrow>int"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    31
  "height_Lam \<equiv> \<lambda>_ _ n. n+1"
19752
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 :: "lam \<Rightarrow> int"
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    34
  "height \<equiv> lam_rec height_Var height_App height_Lam"
19752
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
text {* show that height is a function *}
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    37
lemma fin_supp_height:
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    38
  shows "finite ((supp height_Var)::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    39
  and   "finite ((supp height_App)::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    40
  and   "finite ((supp height_Lam)::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    41
  by (finite_guess add: height_Var_def height_App_def height_Lam_def perm_int_def fs_name1)+
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    42
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    43
lemma fcb_height_Lam:
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    44
  assumes fr: "a\<sharp>height_Lam"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    45
  shows "a\<sharp>height_Lam a t n"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    46
apply(simp add: height_Lam_def perm_int_def fresh_def supp_int)
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    47
done
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    48
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    49
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
    50
lemma height_Var:
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    51
  shows "height (Var c) = 1"
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    52
apply(simp add: height_def)
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    53
apply(simp add: lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_height, OF fcb_height_Lam])
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    54
apply(simp add: height_Var_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
lemma height_App:
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    58
  shows "height (App t1 t2) = (max (height t1) (height t2))+1"
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    59
apply(simp add: height_def)
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    60
apply(simp add: lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_height, OF fcb_height_Lam])
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    61
apply(simp add: height_App_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_Lam:
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    65
  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
    66
apply(simp add: height_def)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    67
apply(rule trans)
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    68
apply(rule lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_height, OF fcb_height_Lam])
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    69
apply(assumption)
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    70
apply(fresh_guess add: height_Var_def perm_int_def)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    71
apply(fresh_guess add: height_App_def perm_int_def)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    72
apply(fresh_guess add: height_Lam_def perm_int_def)
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    73
apply(simp add: height_Lam_def)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    74
done
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    75
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    76
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
    77
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
    78
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    79
text {* define capture-avoiding substitution *}
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    80
constdefs 
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    81
  subst_Var :: "name \<Rightarrow> lam \<Rightarrow> name \<Rightarrow> lam"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    82
  "subst_Var x t' \<equiv> \<lambda>y. (if y=x then t' else (Var y))"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    83
  
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    84
  subst_App :: "name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    85
  "subst_App x t' \<equiv> \<lambda>_ _ r1 r2. App r1 r2"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    86
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    87
  subst_Lam :: "name \<Rightarrow> lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    88
  "subst_Lam x t' \<equiv> \<lambda>a _ r. Lam [a].r"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    89
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    90
  subst_lam :: "name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam" 
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
    91
  "subst_lam x t' \<equiv> lam_rec (subst_Var x t') (subst_App x t') (subst_Lam x t')"
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    92
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    93
lemma supports_subst_Var:
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    94
  shows "((supp (x,t))::name set) supports (subst_Var x t)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    95
apply(supports_simp add: subst_Var_def)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    96
apply(rule impI)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    97
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    98
apply(perm_simp)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    99
done
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   100
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   101
lemma fin_supp_subst:
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   102
  shows "finite ((supp (subst_Var x t))::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   103
  and   "finite ((supp (subst_App x t))::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   104
  and   "finite ((supp (subst_Lam x t))::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   105
proof -
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   106
  case goal1
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   107
  have f: "finite ((supp (x,t))::name set)" by (simp add: fs_name1)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   108
  then have "supp (subst_Var x t) \<subseteq> ((supp (x,t))::name set)"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   109
    using supp_is_subset[OF supports_subst_Var] by simp
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   110
  then show "finite ((supp (subst_Var x t))::name set)" using f by (simp add: finite_subset)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   111
qed (finite_guess add: subst_App_def subst_Lam_def fs_name1)+
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   112
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   113
lemma fcb_subst_Lam:
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   114
  assumes fr: "a\<sharp>(subst_Lam y t')" 
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   115
  shows "a\<sharp>(subst_Lam y t') a t r"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   116
  by (simp add: subst_Lam_def abs_fresh)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   117
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   118
syntax 
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   119
 subst_lam_syn :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   120
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   121
translations 
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   122
  "t1[y::=t2]" \<rightleftharpoons> "subst_lam y t2 t1"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   123
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   124
lemma subst_lam[simp]:
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   125
  shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   126
  and   "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   127
  and   "\<lbrakk>a\<sharp>y; a\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [a].t)[y::=t'] = Lam [a].(t[y::=t'])"
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   128
apply(unfold subst_lam_def)
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
   129
apply(simp only: lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_subst, OF fcb_subst_Lam])
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
   130
apply(simp add: subst_Var_def)
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
   131
apply(simp only: lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_subst, OF fcb_subst_Lam])
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   132
apply(simp only: subst_App_def)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   133
apply(rule trans)
20398
1db76dd407bb modified to use the characteristic equations
urbanc
parents: 20396
diff changeset
   134
apply(rule lam.recs[where P="\<lambda>_. True", simplified, OF fin_supp_subst, OF fcb_subst_Lam])
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   135
apply(assumption)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   136
apply(rule supports_fresh, rule supports_subst_Var, simp add: fs_name1, simp add: fresh_def supp_prod)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   137
apply(fresh_guess add: fresh_prod subst_App_def fs_name1)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   138
apply(fresh_guess add: fresh_prod subst_Lam_def fs_name1)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   139
apply(simp add: subst_Lam_def)
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
   140
done
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   141
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   142
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
   143
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   144
lemma height_ge_one: 
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   145
  shows "1 \<le> (height e)"
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   146
  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
   147
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   148
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
   149
        here entirely by using functions *}
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   150
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   151
theorem height_subst:
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   152
  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
   153
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
   154
  case (Var y)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   155
  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
   156
  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
   157
next
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   158
  case (Lam y e1)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   159
  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
   160
  moreover
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   161
  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
   162
  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
   163
next    
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   164
  case (App e1 e2)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   165
  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
   166
    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
   167
  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
   168
qed
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   169
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   170
end
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
   171