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