thys/Recs.thy
author ibm-PC\ibm <xingyuanzhang@126.com>
Fri, 12 Sep 2014 00:47:15 +0800
changeset 24 77daf1b85cf0
parent 20 e04123f4bacc
permissions -rwxr-xr-x
new change
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
theory Recs
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
imports Main Fact 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
        "~~/src/HOL/Number_Theory/Primes" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
        "~~/src/HOL/Library/Nat_Bijection"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
        "~~/src/HOL/Library/Discrete"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
begin
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
declare One_nat_def[simp del]
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
(*
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
  some definitions from 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
    A Course in Formal Languages, Automata and Groups
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
    I M Chiswell 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
  and
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
    Lecture on undecidability
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
    Michael M. Wolf 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
*)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
lemma if_zero_one [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
  "(if P then 1 else 0) = (0::nat) \<longleftrightarrow> \<not> P"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
  "(0::nat) < (if P then 1 else 0) = P"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
  "(if P then 0 else 1) = (if \<not>P then 1 else (0::nat))"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
by (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
lemma nth:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
  "(x # xs) ! 0 = x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
  "(x # y # xs) ! 1 = y"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  "(x # y # z # xs) ! 2 = z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
  "(x # y # z # u # xs) ! 3 = u"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
by (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
section {* Some auxiliary lemmas about @{text "\<Sum>"} and @{text "\<Prod>"} *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
lemma setprod_atMost_Suc[simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
  "(\<Prod>i \<le> Suc n. f i) = (\<Prod>i \<le> n. f i) * f(Suc n)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
by(simp add:atMost_Suc mult_ac)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
lemma setprod_lessThan_Suc[simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
  "(\<Prod>i < Suc n. f i) = (\<Prod>i < n. f i) * f n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
by (simp add:lessThan_Suc mult_ac)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
lemma setsum_add_nat_ivl2: "n \<le> p  \<Longrightarrow>
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
  setsum f {..<n} + setsum f {n..p} = setsum f {..p::nat}"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
apply(subst setsum_Un_disjoint[symmetric])
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
apply(auto simp add: ivl_disj_un_one)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
lemma setsum_eq_zero [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
  shows "(\<Sum>i < n. f i) = 0 \<longleftrightarrow> (\<forall>i < n. f i = 0)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
        "(\<Sum>i \<le> n. f i) = 0 \<longleftrightarrow> (\<forall>i \<le> n. f i = 0)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
by (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
lemma setprod_eq_zero [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
  shows "(\<Prod>i < n. f i) = 0 \<longleftrightarrow> (\<exists>i < n. f i = 0)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
        "(\<Prod>i \<le> n. f i) = 0 \<longleftrightarrow> (\<exists>i \<le> n. f i = 0)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
by (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
lemma setsum_one_less:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
  fixes n::nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
  assumes "\<forall>i < n. f i \<le> 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
  shows "(\<Sum>i < n. f i) \<le> n"  
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
using assms
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
lemma setsum_one_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
  fixes n::nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
  assumes "\<forall>i \<le> n. f i \<le> 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
  shows "(\<Sum>i \<le> n. f i) \<le> Suc n"  
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
using assms
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
lemma setsum_eq_one_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
  fixes n::nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  assumes "\<forall>i \<le> n. f i = 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
  shows "(\<Sum>i \<le> n. f i) = Suc n"  
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
using assms
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
lemma setsum_least_eq:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
  assumes h0: "p \<le> n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  assumes h1: "\<forall>i \<in> {..<p}. f i = 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
  assumes h2: "\<forall>i \<in> {p..n}. f i = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
  shows "(\<Sum>i \<le> n. f i) = p"  
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
proof -
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
  have eq_p: "(\<Sum>i \<in> {..<p}. f i) = p" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    93
    using h1 by (induct p) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    94
  have eq_zero: "(\<Sum>i \<in> {p..n}. f i) = 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    95
    using h2 by auto
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    96
  have "(\<Sum>i \<le> n. f i) = (\<Sum>i \<in> {..<p}. f i) + (\<Sum>i \<in> {p..n}. f i)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    97
    using h0 by (simp add: setsum_add_nat_ivl2) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    98
  also have "... = (\<Sum>i \<in> {..<p}. f i)" using eq_zero by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  finally show "(\<Sum>i \<le> n. f i) = p" using eq_p by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
qed
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   101
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   102
lemma nat_mult_le_one:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
  fixes m n::nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
  assumes "m \<le> 1" "n \<le> 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
  shows "m * n \<le> 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
using assms by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
lemma setprod_one_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
  assumes "\<forall>i \<le> n. f i \<le> 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
  shows "(\<Prod>i \<le> n. f i) \<le> 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
using assms 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   113
by (induct n) (auto intro: nat_mult_le_one)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
lemma setprod_greater_zero:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
  assumes "\<forall>i \<le> n. f i \<ge> 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
  shows "(\<Prod>i \<le> n. f i) \<ge> 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   119
using assms by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   120
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   121
lemma setprod_eq_one:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   122
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
  assumes "\<forall>i \<le> n. f i = Suc 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  shows "(\<Prod>i \<le> n. f i) = Suc 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
using assms by (induct n) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
lemma setsum_cut_off_less:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   128
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
  assumes h1: "m \<le> n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
  and     h2: "\<forall>i \<in> {m..<n}. f i = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
  shows "(\<Sum>i < n. f i) = (\<Sum>i < m. f i)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
proof -
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
  have eq_zero: "(\<Sum>i \<in> {m..<n}. f i) = 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   134
    using h2 by auto
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   135
  have "(\<Sum>i < n. f i) = (\<Sum>i \<in> {..<m}. f i) + (\<Sum>i \<in> {m..<n}. f i)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   136
    using h1 by (metis atLeast0LessThan le0 setsum_add_nat_ivl) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   137
  also have "... = (\<Sum>i \<in> {..<m}. f i)" using eq_zero by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
  finally show "(\<Sum>i < n. f i) = (\<Sum>i < m. f i)" by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
qed
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
lemma setsum_cut_off_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  fixes f::"nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
  assumes h1: "m \<le> n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
  and     h2: "\<forall>i \<in> {m..n}. f i = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   145
  shows "(\<Sum>i \<le> n. f i) = (\<Sum>i < m. f i)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
proof -
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
  have eq_zero: "(\<Sum>i \<in> {m..n}. f i) = 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
    using h2 by auto
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
  have "(\<Sum>i \<le> n. f i) = (\<Sum>i \<in> {..<m}. f i) + (\<Sum>i \<in> {m..n}. f i)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   150
    using h1 by (simp add: setsum_add_nat_ivl2)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
  also have "... = (\<Sum>i \<in> {..<m}. f i)" using eq_zero by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
  finally show "(\<Sum>i \<le> n. f i) = (\<Sum>i < m. f i)" by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
qed
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   155
lemma setprod_one [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   156
  fixes n::nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   157
  shows "(\<Prod>i < n. Suc 0) = Suc 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   158
        "(\<Prod>i \<le> n. Suc 0) = Suc 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   159
by (induct n) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   160
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   161
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   162
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   163
section {* Recursive Functions *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
datatype recf =  Z
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
              |  S
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
              |  Id nat nat
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   168
              |  Cn nat recf "recf list"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   169
              |  Pr nat recf recf
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
              |  Mn nat recf 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
fun arity :: "recf \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
  where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   174
  "arity Z = 1" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   175
| "arity S = 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
| "arity (Id m n) = m"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
| "arity (Cn n f gs) = n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
| "arity (Pr n f g) = Suc n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   179
| "arity (Mn n f) = n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
text {* Abbreviations for calculating the arity of the constructors *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
abbreviation
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  "CN f gs \<equiv> Cn (arity (hd gs)) f gs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
abbreviation
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
  "PR f g \<equiv> Pr (arity f) f g"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
abbreviation
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
  "MN f \<equiv> Mn (arity f - 1) f"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
text {* the evaluation function and termination relation *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
fun rec_eval :: "recf \<Rightarrow> nat list \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
  "rec_eval Z xs = 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
| "rec_eval S xs = Suc (xs ! 0)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   198
| "rec_eval (Id m n) xs = xs ! n" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
| "rec_eval (Cn n f gs) xs = rec_eval f (map (\<lambda>x. rec_eval x xs) gs)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
| "rec_eval (Pr n f g) (0 # xs) = rec_eval f xs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
| "rec_eval (Pr n f g) (Suc x # xs) = 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
     rec_eval g (x # (rec_eval (Pr n f g) (x # xs)) # xs)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
| "rec_eval (Mn n f) xs = (LEAST x. rec_eval f (x # xs) = 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
inductive 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
  terminates :: "recf \<Rightarrow> nat list \<Rightarrow> bool"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
  termi_z: "terminates Z [n]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
| termi_s: "terminates S [n]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
| termi_id: "\<lbrakk>n < m; length xs = m\<rbrakk> \<Longrightarrow> terminates (Id m n) xs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
| termi_cn: "\<lbrakk>terminates f (map (\<lambda>g. rec_eval g xs) gs); 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
              \<forall>g \<in> set gs. terminates g xs; length xs = n\<rbrakk> \<Longrightarrow> terminates (Cn n f gs) xs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
| termi_pr: "\<lbrakk>\<forall> y < x. terminates g (y # (rec_eval (Pr n f g) (y # xs) # xs));
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
              terminates f xs;
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
              length xs = n\<rbrakk> 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
              \<Longrightarrow> terminates (Pr n f g) (x # xs)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
| termi_mn: "\<lbrakk>length xs = n; terminates f (r # xs); 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
              rec_eval f (r # xs) = 0;
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
              \<forall> i < r. terminates f (i # xs) \<and> rec_eval f (i # xs) > 0\<rbrakk> \<Longrightarrow> terminates (Mn n f) xs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   222
section {* Arithmetic Functions *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
text {*
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   225
  @{text "constn n"} is the recursive function which generates
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   226
  the natural number @{text "n"}. *}
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   227
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
fun constn :: "nat \<Rightarrow> recf"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
  where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
  "constn 0 = Z"  |
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
  "constn (Suc n) = CN S [constn n]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   234
  "rec_swap f = CN f [Id 2 1, Id 2 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
  "rec_add = PR (Id 1 0) (CN S [Id 3 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
  "rec_mult = PR Z (CN rec_add [Id 3 1, Id 3 2])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
  "rec_power = rec_swap (PR (constn 1) (CN rec_mult [Id 3 1, Id 3 2]))"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   245
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
  "rec_fact_aux = PR (constn 1) (CN rec_mult [CN S [Id 3 0], Id 3 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
  "rec_fact = CN rec_fact_aux [Id 1 0, Id 1 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   251
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  "rec_pred = CN (PR Z (Id 3 0)) [Id 1 0, Id 1 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
  "rec_minus = rec_swap (PR (Id 1 0) (CN rec_pred [Id 3 1]))"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   256
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   257
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   258
lemma constn_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
  "rec_eval (constn n) xs = n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
by (induct n) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
lemma swap_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   263
  "rec_eval (rec_swap f) [x, y] = rec_eval f [y, x]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
by (simp add: rec_swap_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   265
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   266
lemma add_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   267
  "rec_eval rec_add [x, y] =  x + y"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   268
by (induct x) (simp_all add: rec_add_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   269
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
lemma mult_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   271
  "rec_eval rec_mult [x, y] = x * y"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   272
by (induct x) (simp_all add: rec_mult_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   274
lemma power_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   275
  "rec_eval rec_power [x, y] = x ^ y"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   276
by (induct y) (simp_all add: rec_power_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
lemma fact_aux_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
  "rec_eval rec_fact_aux [x, y] = fact x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
by (induct x) (simp_all add: rec_fact_aux_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
lemma fact_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
  "rec_eval rec_fact [x] = fact x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   284
by (simp add: rec_fact_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   285
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   286
lemma pred_lemma [simp]: 
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   287
  "rec_eval rec_pred [x] = x - 1"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   288
by (induct x) (simp_all add: rec_pred_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   289
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   290
lemma minus_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
  "rec_eval rec_minus [x, y] = x - y"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
by (induct y) (simp_all add: rec_minus_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   293
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   294
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   295
section {* Logical Functions *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   297
text {* 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   298
  The @{text "sign"} function returns 1 when the input argument 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   299
  is greater than @{text "0"}. *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   301
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
  "rec_sign = CN rec_minus [constn 1, CN rec_minus [constn 1, Id 1 0]]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
  "rec_not = CN rec_minus [constn 1, Id 1 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
text {*
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
  @{text "rec_eq"} compares two arguments: returns @{text "1"}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
  if they are equal; @{text "0"} otherwise. *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   311
  "rec_eq = CN rec_minus [CN (constn 1) [Id 2 0], CN rec_add [rec_minus, rec_swap rec_minus]]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   312
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   313
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
  "rec_noteq = CN rec_not [rec_eq]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  "rec_conj = CN rec_sign [rec_mult]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
  "rec_disj = CN rec_sign [rec_add]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
  "rec_imp = CN rec_disj [CN rec_not [Id 2 0], Id 2 1]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   325
text {* @{term "rec_ifz [z, x, y]"} returns @{text x} if @{text z} is zero, and
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   326
  @{text y} otherwise;  @{term "rec_if [z, x, y]"} returns @{text x} if @{text z} 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   327
  is *not* zero, and @{text y} otherwise. *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
  "rec_ifz = PR (Id 2 0) (Id 4 3)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
  "rec_if = CN rec_ifz [CN rec_not [Id 3 0], Id 3 1, Id 3 2]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
lemma sign_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  "rec_eval rec_sign [x] = (if x = 0 then 0 else 1)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
by (simp add: rec_sign_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
lemma not_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
  "rec_eval rec_not [x] = (if x = 0 then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   342
by (simp add: rec_not_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
lemma eq_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
  "rec_eval rec_eq [x, y] = (if x = y then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
by (simp add: rec_eq_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
lemma noteq_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   349
  "rec_eval rec_noteq [x, y] = (if x \<noteq> y then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
by (simp add: rec_noteq_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
lemma conj_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
  "rec_eval rec_conj [x, y] = (if x = 0 \<or> y = 0 then 0 else 1)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
by (simp add: rec_conj_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
lemma disj_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
  "rec_eval rec_disj [x, y] = (if x = 0 \<and> y = 0 then 0 else 1)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
by (simp add: rec_disj_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
lemma imp_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
  "rec_eval rec_imp [x, y] = (if 0 < x \<and> y = 0 then 0 else 1)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
by (simp add: rec_imp_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
lemma ifz_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   365
  "rec_eval rec_ifz [z, x, y] = (if z = 0 then x else y)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
by (case_tac z) (simp_all add: rec_ifz_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
lemma if_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
  "rec_eval rec_if [z, x, y] = (if 0 < z then x else y)" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
by (simp add: rec_if_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   372
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   373
section {* The Less and Less-Equal Relations *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
text {*
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
  @{text "rec_less"} compares two arguments and returns @{text "1"} if
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   377
  the first is less than the second; otherwise it returns @{text "0"}. *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
  "rec_less = CN rec_sign [rec_swap rec_minus]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
  "rec_le = CN rec_disj [rec_less, rec_eq]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
lemma less_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
  "rec_eval rec_less [x, y] = (if x < y then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
by (simp add: rec_less_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
lemma le_lemma [simp]: 
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   390
  "rec_eval rec_le [x, y] = (if x \<le> y then 1 else 0)"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   391
by (simp add: rec_le_def)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
section {* Summation and Product Functions *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
  "rec_sigma1 f = PR (CN f [CN Z [Id 1 0], Id 1 0]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
                     (CN rec_add [Id 3 1, CN f [CN S [Id 3 0], Id 3 2]])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
  "rec_sigma2 f = PR (CN f [CN Z [Id 2 0], Id 2 0, Id 2 1]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
                     (CN rec_add [Id 4 1, CN f [CN S [Id 4 0], Id 4 2, Id 4 3]])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
  "rec_accum1 f = PR (CN f [CN Z [Id 1 0], Id 1 0]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
                     (CN rec_mult [Id 3 1, CN f [CN S [Id 3 0], Id 3 2]])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
  "rec_accum2 f = PR (CN f [CN Z [Id 2 0], Id 2 0, Id 2 1]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   410
                     (CN rec_mult [Id 4 1, CN f [CN S [Id 4 0], Id 4 2, Id 4 3]])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
  "rec_accum3 f = PR (CN f [CN Z [Id 3 0], Id 3 0, Id 3 1, Id 3 2]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
                     (CN rec_mult [Id 5 1, CN f [CN S [Id 5 0], Id 5 2, Id 5 3, Id 5 4]])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
lemma sigma1_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
  shows "rec_eval (rec_sigma1 f) [x, y] = (\<Sum> z \<le> x. rec_eval f [z, y])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
by (induct x) (simp_all add: rec_sigma1_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
lemma sigma2_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
  shows "rec_eval (rec_sigma2 f) [x, y1, y2] = (\<Sum> z \<le> x. rec_eval f  [z, y1, y2])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
by (induct x) (simp_all add: rec_sigma2_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
lemma accum1_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
  shows "rec_eval (rec_accum1 f) [x, y] = (\<Prod> z \<le> x. rec_eval f  [z, y])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   427
by (induct x) (simp_all add: rec_accum1_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
lemma accum2_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
  shows "rec_eval (rec_accum2 f) [x, y1, y2] = (\<Prod> z \<le> x. rec_eval f  [z, y1, y2])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
by (induct x) (simp_all add: rec_accum2_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
lemma accum3_lemma [simp]: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
  shows "rec_eval (rec_accum3 f) [x, y1, y2, y3] = (\<Prod> z \<le> x. (rec_eval f)  [z, y1, y2, y3])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
by (induct x) (simp_all add: rec_accum3_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
section {* Bounded Quantifiers *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   440
text {* Instead of defining quantifiers taking an aritrary 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   441
  list of arguments, we define the simpler quantifiers taking
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   442
  one, two and three extra arguments besides the argument
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   443
  that is quantified over. *}
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   444
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
  "rec_all1 f = CN rec_sign [rec_accum1 f]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
  "rec_all2 f = CN rec_sign [rec_accum2 f]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
  "rec_all3 f = CN rec_sign [rec_accum3 f]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   454
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
  "rec_all1_less f = (let cond1 = CN rec_eq [Id 3 0, Id 3 1] in
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
                      let cond2 = CN f [Id 3 0, Id 3 2] 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
                      in CN (rec_all2 (CN rec_disj [cond1, cond2])) [Id 2 0, Id 2 0, Id 2 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
  "rec_all2_less f = (let cond1 = CN rec_eq [Id 4 0, Id 4 1] in 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
                      let cond2 = CN f [Id 4 0, Id 4 2, Id 4 3] in 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
                      CN (rec_all3 (CN rec_disj [cond1, cond2])) [Id 3 0, Id 3 0, Id 3 1, Id 3 2])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   463
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
  "rec_ex1 f = CN rec_sign [rec_sigma1 f]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   468
  "rec_ex2 f = CN rec_sign [rec_sigma2 f]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   469
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
lemma ex1_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
 "rec_eval (rec_ex1 f) [x, y] = (if (\<exists>z \<le> x. 0 < rec_eval f [z, y]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
by (simp add: rec_ex1_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
lemma ex2_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
 "rec_eval (rec_ex2 f) [x, y1, y2] = (if (\<exists>z \<le> x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
by (simp add: rec_ex2_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   478
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   479
lemma all1_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
 "rec_eval (rec_all1 f) [x, y] = (if (\<forall>z \<le> x. 0 < rec_eval f [z, y]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
by (simp add: rec_all1_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
lemma all2_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   484
 "rec_eval (rec_all2 f) [x, y1, y2] = (if (\<forall>z \<le> x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   485
by (simp add: rec_all2_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   487
lemma all3_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
 "rec_eval (rec_all3 f) [x, y1, y2, y3] = (if (\<forall>z \<le> x. 0 < rec_eval f [z, y1, y2, y3]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
by (simp add: rec_all3_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
lemma all1_less_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   492
  "rec_eval (rec_all1_less f) [x, y] = (if (\<forall>z < x. 0 < rec_eval f [z, y]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   493
apply(auto simp add: Let_def rec_all1_less_def)
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   494
apply(metis nat_less_le)+
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   496
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   497
lemma all2_less_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
  "rec_eval (rec_all2_less f) [x, y1, y2] = (if (\<forall>z < x. 0 < rec_eval f [z, y1, y2]) then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   499
apply(auto simp add: Let_def rec_all2_less_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   500
apply(metis nat_less_le)+
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   501
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   503
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   504
section {* Natural Number Division *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   506
definition
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   507
  "rec_div = (let lhs = CN S [Id 3 0] in
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   508
              let rhs = CN rec_mult [Id 3 2, CN S [Id 3 1]] in
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   509
              let cond = CN rec_eq [lhs, rhs] in
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   510
              let if_stmt = CN rec_if [cond, CN S [Id 3 1], Id 3 1]
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   511
              in PR Z if_stmt)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   513
fun Div where
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   514
  "Div x 0 = 0"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   515
| "Div x (Suc y) = (if (Suc y = x * (Suc (Div x y))) then Suc (Div x y) else Div x y)"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   516
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   517
lemma Div0:
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   518
  shows "Div 0 y = 0"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   519
by (induct y) (auto)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   521
lemma Div1:
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   522
  "x * (Div x y) \<le> y"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
by (induct y) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   525
lemma Div2: 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   526
  "x * (Div x y) + y mod x = y"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   527
by (induct y) (auto simp add: mod_Suc)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   528
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   529
lemma Div3:
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   530
  "x * (Div x y) = y - y mod x"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   531
using Div2[of x y] by (auto)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   533
lemma Div4:
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
  assumes h: "0 < x"
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   535
  shows "y < x + x * Div x y"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
proof -
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
  have "x - (y mod x) > 0" using mod_less_divisor assms by auto
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   538
  then have "y < y + (x - (y mod x))" by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
  then have "y < x + (y - (y mod x))" by simp
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   540
  then show "y < x + x * (Div x y)" by (simp add: Div3) 
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   541
qed
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   542
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   543
lemma Div_div: 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   544
  shows "Div x y = y div x"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
apply(case_tac "x = 0")
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   546
apply(simp add: Div0)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
apply(subst split_div_lemma[symmetric])
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   548
apply(auto intro: Div1 Div4)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   549
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   550
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   551
lemma Div_rec_div:
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   552
  shows "rec_eval rec_div [y, x] = Div x y"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   553
by (induct y) (simp_all add: rec_div_def)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   555
lemma div_lemma [simp]:
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   556
  shows "rec_eval rec_div [y, x] = y div x"
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   557
by (simp add: Div_div Div_rec_div)
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   558
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   559
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   560
section {* Iteration *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   561
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   562
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   563
   "rec_iter f = PR (Id 1 0) (CN f [Id 3 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   564
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
fun Iter where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   566
  "Iter f 0 = id"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
| "Iter f (Suc n) = f \<circ> (Iter f n)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   569
lemma Iter_comm:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   570
  "(Iter f n) (f x) = f ((Iter f n) x)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
by (induct n) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   573
lemma iter_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   574
  "rec_eval (rec_iter f) [n, x] =  Iter (\<lambda>x. rec_eval f [x]) n x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   575
by (induct n) (simp_all add: rec_iter_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   576
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   577
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   578
section {* Bounded Maximisation *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   580
text {* Computes the greatest number below a bound @{text n}
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   581
  such that a predicate holds. If such a maximum does not exist,
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   582
  then @{text 0} is returned. *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
fun BMax_rec where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
  "BMax_rec R 0 = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
| "BMax_rec R (Suc n) = (if R (Suc n) then (Suc n) else BMax_rec R n)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
  BMax_set :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
where 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
  "BMax_set R x = Max ({z. z \<le> x \<and> R z} \<union> {0})"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   593
lemma BMax_rec_eq1:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
 "BMax_rec R x = (GREATEST z. (R z \<and> z \<le> x) \<or> z = 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   595
apply(induct x)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   596
apply(auto intro: Greatest_equality Greatest_equality[symmetric])
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   597
apply(simp add: le_Suc_eq)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   598
by metis
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   599
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   600
lemma BMax_rec_eq2:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   601
  "BMax_rec R x = Max ({z. z \<le> x \<and> R z} \<union> {0})"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   602
apply(induct x)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   603
apply(auto intro: Max_eqI Max_eqI[symmetric])
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   604
apply(simp add: le_Suc_eq)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   605
by metis
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   607
lemma BMax_rec_eq3:
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   608
  "BMax_rec R x = Max (Set.filter R {..x} \<union> {0})"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
by (simp add: BMax_rec_eq2 Set.filter_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   610
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
  "rec_max1 f = PR Z (CN rec_ifz [CN f [CN S [Id 3 0], Id 3 2], CN S [Id 3 0], Id 3 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
lemma max1_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
  "rec_eval (rec_max1 f) [x, y] = BMax_rec (\<lambda>u. rec_eval f [u, y] = 0) x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
by (induct x) (simp_all add: rec_max1_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
  "rec_max2 f = PR Z (CN rec_ifz [CN f [CN S [Id 4 0], Id 4 2, Id 4 3], CN S [Id 4 0], Id 4 1])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   621
lemma max2_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   622
  "rec_eval (rec_max2 f) [x, y1, y2] = BMax_rec (\<lambda>u. rec_eval f [u, y1, y2] = 0) x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   623
by (induct x) (simp_all add: rec_max2_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   624
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   625
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   626
section {* Encodings using Cantor's Pairing Function *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   628
text {* We use Cantor's pairing function from the theory 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   629
  Nat_Bijection. However, we need to prove that the formulation 
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   630
  of the decoding function there is recursive. For this we first 
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
  prove that we can extract the maximal triangle number 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   632
  using @{term prod_decode}.
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
*}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   635
abbreviation Max_triangle_aux where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   636
  "Max_triangle_aux k z \<equiv> fst (prod_decode_aux k z) + snd (prod_decode_aux k z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   637
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   638
abbreviation Max_triangle where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   639
  "Max_triangle z \<equiv> Max_triangle_aux 0 z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   640
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   641
abbreviation
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   642
  "pdec1 z \<equiv> fst (prod_decode z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   643
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   644
abbreviation
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   645
  "pdec2 z \<equiv> snd (prod_decode z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   646
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   647
abbreviation 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   648
  "penc m n \<equiv> prod_encode (m, n)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   649
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   650
lemma fst_prod_decode: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   651
  "pdec1 z = z - triangle (Max_triangle z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   652
by (subst (3) prod_decode_inverse[symmetric]) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   653
   (simp add: prod_encode_def prod_decode_def split: prod.split)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   654
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   655
lemma snd_prod_decode: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   656
  "pdec2 z = Max_triangle z - pdec1 z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   657
by (simp only: prod_decode_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   658
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
lemma le_triangle:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
  "m \<le> triangle (n + m)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
by (induct_tac m) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   663
lemma Max_triangle_triangle_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   664
  "triangle (Max_triangle z) \<le> z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   665
by (subst (9) prod_decode_inverse[symmetric])
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
   (simp add: prod_decode_def prod_encode_def split: prod.split)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
lemma Max_triangle_le: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
  "Max_triangle z \<le> z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
proof -
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
  have "Max_triangle z \<le> triangle (Max_triangle z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
    using le_triangle[of _ 0, simplified] by simp
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
  also have "... \<le> z" by (rule Max_triangle_triangle_le)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
  finally show "Max_triangle z \<le> z" .
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   675
qed
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   676
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   677
lemma w_aux: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   678
  "Max_triangle (triangle k + m) = Max_triangle_aux k m"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
by (simp add: prod_decode_def[symmetric] prod_decode_triangle_add)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   680
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   681
lemma y_aux: "y \<le> Max_triangle_aux y k"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   682
apply(induct k arbitrary: y rule: nat_less_induct)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
apply(subst (1 2) prod_decode_aux.simps)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   684
apply(simp)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   685
apply(rule impI)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   686
apply(drule_tac x="n - Suc y" in spec)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   687
apply(drule mp)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   688
apply(auto)[1]
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   689
apply(drule_tac x="Suc y" in spec)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   690
apply(erule Suc_leD)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   691
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   693
lemma Max_triangle_greatest: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   694
  "Max_triangle z = (GREATEST k. (triangle k \<le> z \<and> k \<le> z) \<or> k = 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   695
apply(rule Greatest_equality[symmetric])
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   696
apply(rule disjI1)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   697
apply(rule conjI)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   698
apply(rule Max_triangle_triangle_le)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
apply(rule Max_triangle_le)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   700
apply(erule disjE)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   701
apply(erule conjE)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   702
apply(subst (asm) (1) le_iff_add)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   703
apply(erule exE)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   704
apply(clarify)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   705
apply(simp only: w_aux)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   706
apply(rule y_aux)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   707
apply(simp)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   708
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   709
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   710
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   711
definition 
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   712
  "rec_triangle = CN rec_div [CN rec_mult [Id 1 0, S], constn 2]"
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   713
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   714
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   715
  "rec_max_triangle = 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   716
       (let cond = CN rec_not [CN rec_le [CN rec_triangle [Id 2 0], Id 2 1]] in
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   717
        CN (rec_max1 cond) [Id 1 0, Id 1 0])"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   718
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   720
lemma triangle_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   721
  "rec_eval rec_triangle [x] = triangle x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   722
by (simp add: rec_triangle_def triangle_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   723
 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
lemma max_triangle_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   725
  "rec_eval rec_max_triangle [x] = Max_triangle x"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   726
by (simp add: Max_triangle_greatest rec_max_triangle_def Let_def BMax_rec_eq1) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   728
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   729
text {* Encodings for Pairs of Natural Numbers *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   730
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   731
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   732
  "rec_penc = CN rec_add [CN rec_triangle [CN rec_add [Id 2 0, Id 2 1]], Id 2 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   733
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   734
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   735
  "rec_pdec1 = CN rec_minus [Id 1 0, CN rec_triangle [CN rec_max_triangle [Id 1 0]]]" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   736
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   737
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
  "rec_pdec2 = CN rec_minus [CN rec_max_triangle [Id 1 0], CN rec_pdec1 [Id 1 0]]" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   740
lemma pdec1_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
  "rec_eval rec_pdec1 [z] = pdec1 z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   742
by (simp add: rec_pdec1_def fst_prod_decode)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   743
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
lemma pdec2_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   745
  "rec_eval rec_pdec2 [z] = pdec2 z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   746
by (simp add: rec_pdec2_def snd_prod_decode)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   747
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   748
lemma penc_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   749
  "rec_eval rec_penc [m, n] = penc m n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   750
by (simp add: rec_penc_def prod_encode_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   751
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   752
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   753
text {* Encodings of Lists of Natural Numbers *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   754
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   755
fun 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   756
  lenc :: "nat list \<Rightarrow> nat" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   757
where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   758
  "lenc [] = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   759
| "lenc (x # xs) = penc (Suc x) (lenc xs)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
fun
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
  ldec :: "nat \<Rightarrow> nat \<Rightarrow> nat"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   764
  "ldec z 0 = (pdec1 z) - 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   765
| "ldec z (Suc n) = ldec (pdec2 z) n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   766
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   767
lemma pdec_zero_simps [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   768
  "pdec1 0 = 0" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   769
  "pdec2 0 = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   770
by (simp_all add: prod_decode_def prod_decode_aux.simps)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   771
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   772
lemma ldec_zero:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   773
  "ldec 0 n = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   774
by (induct n) (simp_all add: prod_decode_def prod_decode_aux.simps)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   775
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   776
lemma list_encode_inverse: 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   777
  "ldec (lenc xs) n = (if n < length xs then xs ! n else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   778
by (induct xs arbitrary: n rule: lenc.induct) 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   779
   (auto simp add: ldec_zero nth_Cons split: nat.splits)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   780
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   781
lemma lenc_length_le:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   782
  "length xs \<le> lenc xs"  
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   783
by (induct xs) (simp_all add: prod_encode_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   784
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   785
text {* The membership predicate for an encoded list. *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   786
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   787
fun within :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   788
  "within z 0 = (0 < z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   789
| "within z (Suc n) = within (pdec2 z) n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   790
    
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   791
definition enclen :: "nat \<Rightarrow> nat" where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   792
  "enclen z = BMax_rec (\<lambda>x. within z (x - 1)) z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   793
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   794
lemma within_False [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   795
  "within 0 n = False"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   796
by (induct n) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   797
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   798
lemma within_length [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   799
  "within (lenc xs) s = (s < length xs)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   800
apply(induct s arbitrary: xs)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   801
apply(case_tac xs)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
apply(simp_all add: prod_encode_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   803
apply(case_tac xs)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
apply(simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
20
e04123f4bacc soem more work
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
   807
text {* The length of an encoded list. *}
15
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   809
lemma enclen_length [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810
  "enclen (lenc xs) = length xs"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
unfolding enclen_def
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
apply(simp add: BMax_rec_eq1)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   813
apply(rule Greatest_equality)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   814
apply(auto simp add: lenc_length_le)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   815
done
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   817
lemma enclen_penc [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   818
  "enclen (penc (Suc x) (lenc xs)) = Suc (enclen (lenc xs))"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
by (simp only: lenc.simps[symmetric] enclen_length) (simp)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   821
lemma enclen_zero [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   822
  "enclen 0 = 0"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   823
by (simp add: enclen_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   824
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   825
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
text {* Recursive Definitions for List Encodings *}
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   828
fun 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   829
  rec_lenc :: "recf list \<Rightarrow> recf" 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   830
where
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   831
  "rec_lenc [] = Z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   832
| "rec_lenc (f # fs) = CN rec_penc [CN S [f], rec_lenc fs]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   833
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   834
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   835
  "rec_ldec = CN rec_pred [CN rec_pdec1 [rec_swap (rec_iter rec_pdec2)]]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   836
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   837
definition 
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   838
  "rec_within = CN rec_less [Z, rec_swap (rec_iter rec_pdec2)]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   839
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   840
definition
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   841
  "rec_enclen = CN (rec_max1 (CN rec_not [CN rec_within [Id 2 1, CN rec_pred [Id 2 0]]])) [Id 1 0, Id 1 0]"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   842
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   843
lemma ldec_iter:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   844
  "ldec z n = pdec1 (Iter pdec2 n z) - 1"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   845
by (induct n arbitrary: z) (simp | subst Iter_comm)+
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   846
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   847
lemma within_iter:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   848
  "within z n = (0 < Iter pdec2 n z)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   849
by (induct n arbitrary: z) (simp | subst Iter_comm)+
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   850
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   851
lemma lenc_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
  "rec_eval (rec_lenc fs) xs = lenc (map (\<lambda>f. rec_eval f xs) fs)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
by (induct fs) (simp_all)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
lemma ldec_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   856
  "rec_eval rec_ldec [z, n] = ldec z n"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   857
by (simp add: ldec_iter rec_ldec_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
lemma within_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
  "rec_eval rec_within [z, n] = (if within z n then 1 else 0)"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   861
by (simp add: within_iter rec_within_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
lemma enclen_lemma [simp]:
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
  "rec_eval rec_enclen [z] = enclen z"
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
by (simp add: rec_enclen_def enclen_def)
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   868
end
e3ecf558aef2 recursive function theories / UF_rec still need coding of tapes and programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   869