src/HOL/Nominal/Examples/Height.thy
author berghofe
Wed, 07 Feb 2007 17:44:07 +0100
changeset 22271 51a80e238b29
parent 21555 229c0158de92
child 22418 49e2d9744ae1
permissions -rw-r--r--
Adapted to new inductive definition package.
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
21053
7d0962594902 cleaning up
urbanc
parents: 20398
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
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    15
text {* definition of the height-function *} 
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    16
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    17
consts 
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    18
  height :: "lam \<Rightarrow> int"
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    19
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    20
nominal_primrec
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    21
  "height (Var x) = 1"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    22
  "height (App t1 t2) = (max (height t1) (height t2)) + 1"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    23
  "height (Lam [a].t) = (height t) + 1"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    24
  apply(finite_guess add: perm_int_def)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    25
  apply(rule TrueI)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    26
  apply(simp add: fresh_int)
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    27
  apply(fresh_guess add: perm_int_def)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    28
  done
21087
3e56528a39f7 Adapted to changes in FCBs.
berghofe
parents: 21053
diff changeset
    29
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    30
text {* definition of capture-avoiding substitution *}
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    31
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    32
lemma eq_eqvt:
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    33
  fixes pi::"name prm"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    34
  and   x::"'a::pt_name"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    35
  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    36
  apply(simp add: perm_bool perm_bij)
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    37
  done
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    38
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    39
consts
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    40
  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
    41
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    42
nominal_primrec
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    43
  "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    44
  "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    45
  "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    46
apply(finite_guess add: eq_eqvt perm_if fs_name1)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    47
apply(rule TrueI)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    48
apply(simp add: abs_fresh)
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    49
apply(fresh_guess add: eq_eqvt perm_if fs_name1)+
20396
4d0c33719348 used the recursion combinator for the height and substitution function
urbanc
parents: 19752
diff changeset
    50
done
19752
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
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
    53
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    54
lemma height_ge_one: 
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    55
  shows "1 \<le> (height e)"
21053
7d0962594902 cleaning up
urbanc
parents: 20398
diff changeset
    56
  by (nominal_induct e rule: lam.induct) 
7d0962594902 cleaning up
urbanc
parents: 20398
diff changeset
    57
     (simp | arith)+
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    58
21053
7d0962594902 cleaning up
urbanc
parents: 20398
diff changeset
    59
text {* unlike the proplem suggested by Wang, however, the 
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    60
        theorem is here formulated here entirely by using 
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    61
        functions *}
19752
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    62
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    63
theorem height_subst:
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    64
  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
    65
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
    66
  case (Var y)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    67
  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
    68
  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
    69
next
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    70
  case (Lam y e1)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    71
  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
    72
  moreover
21053
7d0962594902 cleaning up
urbanc
parents: 20398
diff changeset
    73
  have vc: "y\<sharp>x" "y\<sharp>e'" by fact (* usual variable convention *)
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21087
diff changeset
    74
  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
    75
next    
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    76
  case (App e1 e2)
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    77
  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
    78
    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
    79
  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
    80
qed
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    81
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    82
end
18e5aa65fb8b added an example suggested by D. Wang on the PoplMark-mailing list;
urbanc
parents:
diff changeset
    83