src/HOL/Library/Wfrec.thy
changeset 44014 88bd7d74a2c1
child 44259 b922e91dd1d9
equal deleted inserted replaced
44013:5cfc1c36ae97 44014:88bd7d74a2c1
       
     1 (*  Title:      HOL/Library/Wfrec.thy
       
     2     Author:     Tobias Nipkow
       
     3     Author:     Lawrence C Paulson
       
     4     Author:     Konrad Slind
       
     5 *)
       
     6 
       
     7 header {* Well-Founded Recursion Combinator *}
       
     8 
       
     9 theory Wfrec
       
    10 imports Main
       
    11 begin
       
    12 
       
    13 inductive
       
    14   wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool"
       
    15   for R :: "('a * 'a) set"
       
    16   and F :: "('a => 'b) => 'a => 'b"
       
    17 where
       
    18   wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
       
    19             wfrec_rel R F x (F g x)"
       
    20 
       
    21 definition
       
    22   cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" where
       
    23   "cut f r x == (%y. if (y,x):r then f y else undefined)"
       
    24 
       
    25 definition
       
    26   adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" where
       
    27   "adm_wf R F == ALL f g x.
       
    28      (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
       
    29 
       
    30 definition
       
    31   wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" where
       
    32   "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
       
    33 
       
    34 lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"
       
    35 by (simp add: fun_eq_iff cut_def)
       
    36 
       
    37 lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)"
       
    38 by (simp add: cut_def)
       
    39 
       
    40 text{*Inductive characterization of wfrec combinator; for details see:
       
    41 John Harrison, "Inductive definitions: automation and application"*}
       
    42 
       
    43 lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y"
       
    44 apply (simp add: adm_wf_def)
       
    45 apply (erule_tac a=x in wf_induct)
       
    46 apply (rule ex1I)
       
    47 apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI)
       
    48 apply (fast dest!: theI')
       
    49 apply (erule wfrec_rel.cases, simp)
       
    50 apply (erule allE, erule allE, erule allE, erule mp)
       
    51 apply (fast intro: the_equality [symmetric])
       
    52 done
       
    53 
       
    54 lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
       
    55 apply (simp add: adm_wf_def)
       
    56 apply (intro strip)
       
    57 apply (rule cuts_eq [THEN iffD2, THEN subst], assumption)
       
    58 apply (rule refl)
       
    59 done
       
    60 
       
    61 lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"
       
    62 apply (simp add: wfrec_def)
       
    63 apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption)
       
    64 apply (rule wfrec_rel.wfrecI)
       
    65 apply (intro strip)
       
    66 apply (erule adm_lemma [THEN wfrec_unique, THEN theI'])
       
    67 done
       
    68 
       
    69 
       
    70 text{** This form avoids giant explosions in proofs.  NOTE USE OF ==*}
       
    71 lemma def_wfrec: "[| f==wfrec r H;  wf(r) |] ==> f(a) = H (cut f r a) a"
       
    72 apply auto
       
    73 apply (blast intro: wfrec)
       
    74 done
       
    75 
       
    76 
       
    77 subsection {* Nitpick setup *}
       
    78 
       
    79 axiomatization wf_wfrec :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
       
    80 
       
    81 definition wf_wfrec' :: "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
       
    82 [nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x"
       
    83 
       
    84 definition wfrec' ::  "('a \<times> 'a \<Rightarrow> bool) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
       
    85 "wfrec' R F x \<equiv> if wf R then wf_wfrec' R F x
       
    86                 else THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
       
    87 
       
    88 setup {*
       
    89   Nitpick_HOL.register_ersatz_global
       
    90     [(@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
       
    91      (@{const_name wfrec}, @{const_name wfrec'})]
       
    92 *}
       
    93 
       
    94 hide_const (open) wf_wfrec wf_wfrec' wfrec'
       
    95 hide_fact (open) wf_wfrec'_def wfrec'_def
       
    96 
       
    97 subsection {* Wellfoundedness of @{text same_fst} *}
       
    98 
       
    99 definition
       
   100  same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
       
   101 where
       
   102     "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
       
   103    --{*For @{text rec_def} declarations where the first n parameters
       
   104        stay unchanged in the recursive call. *}
       
   105 
       
   106 lemma same_fstI [intro!]:
       
   107      "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
       
   108 by (simp add: same_fst_def)
       
   109 
       
   110 lemma wf_same_fst:
       
   111   assumes prem: "(!!x. P x ==> wf(R x))"
       
   112   shows "wf(same_fst P R)"
       
   113 apply (simp cong del: imp_cong add: wf_def same_fst_def)
       
   114 apply (intro strip)
       
   115 apply (rename_tac a b)
       
   116 apply (case_tac "wf (R a)")
       
   117  apply (erule_tac a = b in wf_induct, blast)
       
   118 apply (blast intro: prem)
       
   119 done
       
   120 
       
   121 end