src/HOL/Nominal/Examples/Recursion.thy
author urbanc
Thu, 02 Mar 2006 15:05:09 +0100
changeset 19171 17b952ec5641
child 19651 247ca17caddd
permissions -rw-r--r--
split the files - Iteration.thy contains the big proof of the iteration combinator - Recursion.thy derives from Iteration the recursion combinator - lam_substs.thy contains the examples (size, substitution and parallel substitution)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19171
17b952ec5641 split the files
urbanc
parents:
diff changeset
     1
(* $Id$ *)
17b952ec5641 split the files
urbanc
parents:
diff changeset
     2
17b952ec5641 split the files
urbanc
parents:
diff changeset
     3
theory Recursion
17b952ec5641 split the files
urbanc
parents:
diff changeset
     4
imports "Iteration" 
17b952ec5641 split the files
urbanc
parents:
diff changeset
     5
begin
17b952ec5641 split the files
urbanc
parents:
diff changeset
     6
17b952ec5641 split the files
urbanc
parents:
diff changeset
     7
types 'a f1_ty' = "name\<Rightarrow>('a::pt_name)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
     8
      'a f2_ty' = "lam\<Rightarrow>lam\<Rightarrow>'a\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
     9
      'a f3_ty' = "lam\<Rightarrow>name\<Rightarrow>'a\<Rightarrow>('a::pt_name)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    10
17b952ec5641 split the files
urbanc
parents:
diff changeset
    11
constdefs
17b952ec5641 split the files
urbanc
parents:
diff changeset
    12
  rfun' :: "'a f1_ty' \<Rightarrow> 'a f2_ty' \<Rightarrow> 'a f3_ty' \<Rightarrow> lam \<Rightarrow> (lam\<times>'a::pt_name)" 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    13
  "rfun' f1 f2 f3 t \<equiv> 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    14
      (itfun 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    15
        (\<lambda>a. (Var a,f1 a)) 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    16
        (\<lambda>r1 r2. (App (fst r1) (fst r2), f2 (fst r1) (fst r2) (snd r1) (snd r2))) 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    17
        (\<lambda>a r. (Lam [a].(fst r),f3 (fst r) a (snd r)))
17b952ec5641 split the files
urbanc
parents:
diff changeset
    18
       t)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    19
17b952ec5641 split the files
urbanc
parents:
diff changeset
    20
  rfun :: "'a f1_ty' \<Rightarrow> 'a f2_ty' \<Rightarrow> 'a f3_ty' \<Rightarrow> lam \<Rightarrow> 'a::pt_name" 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    21
  "rfun f1 f2 f3 t \<equiv> snd (rfun' f1 f2 f3 t)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    22
17b952ec5641 split the files
urbanc
parents:
diff changeset
    23
lemma fcb': 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    24
  fixes f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    25
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    26
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    27
  shows  "\<exists>a.  a \<sharp> (\<lambda>a r. (Lam [a].fst r, f3 (fst r) a (snd r))) \<and>
17b952ec5641 split the files
urbanc
parents:
diff changeset
    28
               (\<forall>y.  a \<sharp> (Lam [a].fst y, f3 (fst y) a (snd y)))"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    29
using c f
17b952ec5641 split the files
urbanc
parents:
diff changeset
    30
apply(auto)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    31
apply(rule_tac x="a" in exI)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    32
apply(auto simp add: abs_fresh fresh_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    33
apply(rule_tac S="supp f3" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    34
apply(supports_simp add: perm_fst perm_snd)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    35
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    36
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    37
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
    38
17b952ec5641 split the files
urbanc
parents:
diff changeset
    39
lemma fsupp':
17b952ec5641 split the files
urbanc
parents:
diff changeset
    40
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    41
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    42
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    43
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    44
  shows "finite((supp
17b952ec5641 split the files
urbanc
parents:
diff changeset
    45
          (\<lambda>a. (Var a, f1 a),
17b952ec5641 split the files
urbanc
parents:
diff changeset
    46
           \<lambda>r1 r2. (App (fst r1) (fst r2), f2 (fst r1) (fst r2) (snd r1) (snd r2)),
17b952ec5641 split the files
urbanc
parents:
diff changeset
    47
           \<lambda>a r. (Lam [a].fst r, f3 (fst r) a (snd r))))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    48
using f by (finite_guess add: perm_fst perm_snd fs_name1 supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    49
17b952ec5641 split the files
urbanc
parents:
diff changeset
    50
lemma rfun'_fst:
17b952ec5641 split the files
urbanc
parents:
diff changeset
    51
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    52
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    53
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    54
  assumes f: "finite ((supp (f1,f2,f3))::name set)" 
17b952ec5641 split the files
urbanc
parents:
diff changeset
    55
  and     c: "(\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y))"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    56
  shows "fst (rfun' f1 f2 f3 t) = t"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    57
apply(rule lam.induct'[of "\<lambda>_. ((supp (f1,f2,f3))::name set)" "\<lambda>_ t. fst (rfun' f1 f2 f3 t) = t"])
17b952ec5641 split the files
urbanc
parents:
diff changeset
    58
apply(fold fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    59
apply(simp add: f)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    60
apply(unfold rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    61
apply(simp only: itfun_Var[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
    62
apply(simp)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    63
apply(simp only: itfun_App[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
    64
apply(simp)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    65
apply(auto)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    66
apply(rule trans)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    67
apply(rule_tac f="fst" in arg_cong)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    68
apply(rule itfun_Lam[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
    69
apply(auto simp add: fresh_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    70
apply(rule_tac S="supp f1" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    71
apply(supports_simp)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    72
apply(subgoal_tac "finite ((supp (f1,f2,f3))::name set)")
17b952ec5641 split the files
urbanc
parents:
diff changeset
    73
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    74
apply(rule f)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    75
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    76
apply(rule_tac S="supp f2" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    77
apply(supports_simp add: perm_fst perm_snd)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    78
apply(subgoal_tac "finite ((supp (f1,f2,f3))::name set)")
17b952ec5641 split the files
urbanc
parents:
diff changeset
    79
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    80
apply(rule f)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    81
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    82
apply(rule_tac S="supp f3" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    83
apply(supports_simp add: perm_fst perm_snd)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    84
apply(subgoal_tac "finite ((supp (f1,f2,f3))::name set)")
17b952ec5641 split the files
urbanc
parents:
diff changeset
    85
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    86
apply(rule f)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    87
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    88
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
    89
17b952ec5641 split the files
urbanc
parents:
diff changeset
    90
lemma rfun'_Var:
17b952ec5641 split the files
urbanc
parents:
diff changeset
    91
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    92
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    93
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    94
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    95
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    96
  shows "rfun' f1 f2 f3 (Var c) = (Var c, f1 c)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
    97
apply(simp add: rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
    98
apply(simp add: itfun_Var[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
    99
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   100
17b952ec5641 split the files
urbanc
parents:
diff changeset
   101
lemma rfun'_App:
17b952ec5641 split the files
urbanc
parents:
diff changeset
   102
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   103
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   104
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   105
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   106
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   107
  shows "rfun' f1 f2 f3 (App t1 t2) = 
17b952ec5641 split the files
urbanc
parents:
diff changeset
   108
                (App t1 t2, f2 t1 t2 (rfun f1 f2 f3 t1) (rfun f1 f2 f3 t2))"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   109
apply(simp add: rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   110
apply(rule trans)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   111
apply(rule itfun_App[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   112
apply(fold rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   113
apply(simp_all add: rfun'_fst[OF f, OF c])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   114
apply(simp_all add: rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   115
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   116
17b952ec5641 split the files
urbanc
parents:
diff changeset
   117
lemma rfun'_Lam:
17b952ec5641 split the files
urbanc
parents:
diff changeset
   118
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   119
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   120
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   121
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   122
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   123
  and     a: "b\<sharp>(f1,f2,f3)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   124
  shows "rfun' f1 f2 f3 (Lam [b].t) = (Lam [b].t, f3 t b (rfun f1 f2 f3 t))"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   125
using a f
17b952ec5641 split the files
urbanc
parents:
diff changeset
   126
apply(simp add: rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   127
apply(rule trans)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   128
apply(rule itfun_Lam[OF fsupp'[OF f],OF fcb'[OF f, OF c]])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   129
apply(auto simp add: fresh_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   130
apply(rule_tac S="supp f1" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   131
apply(supports_simp)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   132
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   133
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   134
apply(rule_tac S="supp f2" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   135
apply(supports_simp add: perm_fst perm_snd)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   136
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   137
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   138
apply(rule_tac S="supp f3" in supports_fresh)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   139
apply(supports_simp add: perm_snd perm_fst)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   140
apply(simp add: supp_prod)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   141
apply(simp add: fresh_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   142
apply(fold rfun'_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   143
apply(simp_all add: rfun'_fst[OF f, OF c])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   144
apply(simp_all add: rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   145
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   146
17b952ec5641 split the files
urbanc
parents:
diff changeset
   147
lemma rfun_Var:
17b952ec5641 split the files
urbanc
parents:
diff changeset
   148
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   149
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   150
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   151
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   152
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   153
  shows "rfun f1 f2 f3 (Var c) = f1 c"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   154
apply(unfold rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   155
apply(simp add: rfun'_Var[OF f, OF c])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   156
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   157
17b952ec5641 split the files
urbanc
parents:
diff changeset
   158
lemma rfun_App:
17b952ec5641 split the files
urbanc
parents:
diff changeset
   159
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   160
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   161
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   162
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   163
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   164
  shows "rfun f1 f2 f3 (App t1 t2) = f2 t1 t2 (rfun f1 f2 f3 t1) (rfun f1 f2 f3 t2)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   165
apply(unfold rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   166
apply(simp add: rfun'_App[OF f, OF c])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   167
apply(simp add: rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   168
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   169
17b952ec5641 split the files
urbanc
parents:
diff changeset
   170
lemma rfun_Lam:
17b952ec5641 split the files
urbanc
parents:
diff changeset
   171
  fixes f1::"('a::pt_name) f1_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   172
  and   f2::"('a::pt_name) f2_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   173
  and   f3::"('a::pt_name) f3_ty'"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   174
  assumes f: "finite ((supp (f1,f2,f3))::name set)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   175
  and     c: "\<exists>(a::name). a\<sharp>f3 \<and> (\<forall>(y::'a::pt_name) t. a\<sharp>f3 t a y)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   176
  and     a: "b\<sharp>(f1,f2,f3)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   177
  shows "rfun f1 f2 f3 (Lam [b].t) = f3 t b (rfun f1 f2 f3 t)"
17b952ec5641 split the files
urbanc
parents:
diff changeset
   178
using a
17b952ec5641 split the files
urbanc
parents:
diff changeset
   179
apply(unfold rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   180
apply(simp add: rfun'_Lam[OF f, OF c])
17b952ec5641 split the files
urbanc
parents:
diff changeset
   181
apply(simp add: rfun_def)
17b952ec5641 split the files
urbanc
parents:
diff changeset
   182
done
17b952ec5641 split the files
urbanc
parents:
diff changeset
   183
17b952ec5641 split the files
urbanc
parents:
diff changeset
   184
end