| author | mengj | 
| Thu, 25 May 2006 08:09:10 +0200 | |
| changeset 19720 | f68f6f958a1d | 
| parent 19602 | e25d1e9a0675 | 
| child 19766 | 031e0dde31f1 | 
| permissions | -rw-r--r-- | 
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 1 | (* ID: $Id$ | 
| 10213 | 2 | Author: Tobias Nipkow | 
| 3 | Copyright 1992 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 6 | header {*Well-founded Recursion*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 7 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 8 | theory Wellfounded_Recursion | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 9 | imports Transitive_Closure | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 10 | begin | 
| 10213 | 11 | |
| 11328 | 12 | consts | 
| 13 |   wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => ('a * 'b) set"
 | |
| 14 | ||
| 15 | inductive "wfrec_rel R F" | |
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 16 | intros | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 17 | wfrecI: "ALL z. (z, x) : R --> (z, g z) : wfrec_rel R F ==> | 
| 11328 | 18 | (x, F g x) : wfrec_rel R F" | 
| 19 | ||
| 10213 | 20 | constdefs | 
| 21 |   wf         :: "('a * 'a)set => bool"
 | |
| 22 | "wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))" | |
| 23 | ||
| 24 |   acyclic :: "('a*'a)set => bool"
 | |
| 25 | "acyclic r == !x. (x,x) ~: r^+" | |
| 26 | ||
| 27 |   cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
 | |
| 28 | "cut f r x == (%y. if (y,x):r then f y else arbitrary)" | |
| 29 | ||
| 11328 | 30 |   adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
 | 
| 31 | "adm_wf R F == ALL f g x. | |
| 32 | (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x" | |
| 10213 | 33 | |
| 11328 | 34 |   wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
 | 
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
11328diff
changeset | 35 | "wfrec R F == %x. THE y. (x, y) : wfrec_rel R (%f x. F (cut f R x) x)" | 
| 10213 | 36 | |
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 37 | axclass wellorder \<subseteq> linorder | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 38 |   wf: "wf {(x,y::'a::ord). x<y}"
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 39 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 40 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 41 | lemma wfUNIVI: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 42 | "(!!P x. (ALL x. (ALL y. (y,x) : r --> P(y)) --> P(x)) ==> P(x)) ==> wf(r)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 43 | by (unfold wf_def, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 44 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 45 | text{*Restriction to domain @{term A}.  
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 46 |   If @{term r} is well-founded over @{term A} then @{term "wf r"}*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 47 | lemma wfI: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 48 | "[| r <= A <*> A; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 49 | !!x P. [| ALL x. (ALL y. (y,x) : r --> P y) --> P x; x:A |] ==> P x |] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 50 | ==> wf r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 51 | by (unfold wf_def, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 52 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 53 | lemma wf_induct: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 54 | "[| wf(r); | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 55 | !!x.[| ALL y. (y,x): r --> P(y) |] ==> P(x) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 56 | |] ==> P(a)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 57 | by (unfold wf_def, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 58 | |
| 18458 | 59 | lemmas wf_induct_rule = wf_induct [rule_format, consumes 1, case_names less, induct set: wf] | 
| 17042 
da5cfaa258f7
moved wf_induct_rule from PreList.thy to Wellfounded_Recursion.thy
 nipkow parents: 
15950diff
changeset | 60 | |
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 61 | lemma wf_not_sym [rule_format]: "wf(r) ==> ALL x. (a,x):r --> (x,a)~:r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 62 | by (erule_tac a=a in wf_induct, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 63 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 64 | (* [| wf r; ~Z ==> (a,x) : r; (x,a) ~: r ==> Z |] ==> Z *) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 65 | lemmas wf_asym = wf_not_sym [elim_format] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 66 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 67 | lemma wf_not_refl [simp]: "wf(r) ==> (a,a) ~: r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 68 | by (blast elim: wf_asym) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 69 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 70 | (* [| wf r; (a,a) ~: r ==> PROP W |] ==> PROP W *) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 71 | lemmas wf_irrefl = wf_not_refl [elim_format] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 72 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 73 | text{*transitive closure of a well-founded relation is well-founded! *}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 74 | lemma wf_trancl: "wf(r) ==> wf(r^+)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 75 | apply (subst wf_def, clarify) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 76 | apply (rule allE, assumption) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 77 |   --{*Retains the universal formula for later use!*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 78 | apply (erule mp) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 79 | apply (erule_tac a = x in wf_induct) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 80 | apply (blast elim: tranclE) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 81 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 82 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 83 | lemma wf_converse_trancl: "wf (r^-1) ==> wf ((r^+)^-1)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 84 | apply (subst trancl_converse [symmetric]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 85 | apply (erule wf_trancl) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 86 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 87 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 88 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 89 | subsubsection{*Minimal-element characterization of well-foundedness*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 90 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 91 | lemma lemma1: "wf r ==> x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 92 | apply (unfold wf_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 93 | apply (drule spec) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 94 | apply (erule mp [THEN spec], blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 95 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 96 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 97 | lemma lemma2: "(ALL Q x. x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q)) ==> wf r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 98 | apply (unfold wf_def, clarify) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 99 | apply (drule_tac x = "{x. ~ P x}" in spec, blast)
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 100 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 101 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 102 | lemma wf_eq_minimal: "wf r = (ALL Q x. x:Q --> (EX z:Q. ALL y. (y,z):r --> y~:Q))" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 103 | by (blast intro!: lemma1 lemma2) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 104 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 105 | subsubsection{*Other simple well-foundedness results*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 106 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 107 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 108 | text{*Well-foundedness of subsets*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 109 | lemma wf_subset: "[| wf(r); p<=r |] ==> wf(p)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 110 | apply (simp (no_asm_use) add: wf_eq_minimal) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 111 | apply fast | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 112 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 113 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 114 | text{*Well-foundedness of the empty relation*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 115 | lemma wf_empty [iff]: "wf({})"
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 116 | by (simp add: wf_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 117 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 118 | text{*Well-foundedness of insert*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 119 | lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 120 | apply (rule iffI) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 121 | apply (blast elim: wf_trancl [THEN wf_irrefl] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 122 | intro: rtrancl_into_trancl1 wf_subset | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 123 | rtrancl_mono [THEN [2] rev_subsetD]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 124 | apply (simp add: wf_eq_minimal, safe) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 125 | apply (rule allE, assumption, erule impE, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 126 | apply (erule bexE) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 127 | apply (rename_tac "a", case_tac "a = x") | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 128 | prefer 2 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 129 | apply blast | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 130 | apply (case_tac "y:Q") | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 131 | prefer 2 apply blast | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 132 | apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE)
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 133 | apply assumption | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 134 | apply (erule_tac V = "ALL Q. (EX x. x : Q) --> ?P Q" in thin_rl) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 135 |   --{*essential for speed*}
 | 
| 15343 | 136 | txt{*Blast with new substOccur fails*}
 | 
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 137 | apply (fast intro: converse_rtrancl_into_rtrancl) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 138 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 139 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 140 | text{*Well-foundedness of image*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 141 | lemma wf_prod_fun_image: "[| wf r; inj f |] ==> wf(prod_fun f f ` r)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 142 | apply (simp only: wf_eq_minimal, clarify) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 143 | apply (case_tac "EX p. f p : Q") | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 144 | apply (erule_tac x = "{p. f p : Q}" in allE)
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 145 | apply (fast dest: inj_onD, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 146 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 147 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 148 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 149 | subsubsection{*Well-Foundedness Results for Unions*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 150 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 151 | text{*Well-foundedness of indexed union with disjoint domains and ranges*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 152 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 153 | lemma wf_UN: "[| ALL i:I. wf(r i); | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 154 |          ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {}  
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 155 | |] ==> wf(UN i:I. r i)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 156 | apply (simp only: wf_eq_minimal, clarify) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 157 | apply (rename_tac A a, case_tac "EX i:I. EX a:A. EX b:A. (b,a) : r i") | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 158 | prefer 2 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 159 | apply force | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 160 | apply clarify | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 161 | apply (drule bspec, assumption) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 162 | apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r i) }" in allE)
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 163 | apply (blast elim!: allE) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 164 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 165 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 166 | lemma wf_Union: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 167 | "[| ALL r:R. wf r; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 168 |      ALL r:R. ALL s:R. r ~= s --> Domain r Int Range s = {}  
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 169 | |] ==> wf(Union R)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 170 | apply (simp add: Union_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 171 | apply (blast intro: wf_UN) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 172 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 173 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 174 | (*Intuition: we find an (R u S)-min element of a nonempty subset A | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 175 | by case distinction. | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 176 | 1. There is a step a -R-> b with a,b : A. | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 177 |      Pick an R-min element z of the (nonempty) set {a:A | EX b:A. a -R-> b}.
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 178 | By definition, there is z':A s.t. z -R-> z'. Because z is R-min in the | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 179 | subset, z' must be R-min in A. Because z' has an R-predecessor, it cannot | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 180 | have an S-successor and is thus S-min in A as well. | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 181 | 2. There is no such step. | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 182 | Pick an S-min element of A. In this case it must be an R-min | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 183 | element of A as well. | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 184 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 185 | *) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 186 | lemma wf_Un: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 187 |      "[| wf r; wf s; Domain r Int Range s = {} |] ==> wf(r Un s)"
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 188 | apply (simp only: wf_eq_minimal, clarify) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 189 | apply (rename_tac A a) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 190 | apply (case_tac "EX a:A. EX b:A. (b,a) : r") | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 191 | prefer 2 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 192 | apply simp | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 193 | apply (drule_tac x=A in spec)+ | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 194 | apply blast | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 195 | apply (erule_tac x="{a. a:A & (EX b:A. (b,a) : r) }" in allE)+
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 196 | apply (blast elim!: allE) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 197 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 198 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 199 | subsubsection {*acyclic*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 200 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 201 | lemma acyclicI: "ALL x. (x, x) ~: r^+ ==> acyclic r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 202 | by (simp add: acyclic_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 203 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 204 | lemma wf_acyclic: "wf r ==> acyclic r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 205 | apply (simp add: acyclic_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 206 | apply (blast elim: wf_trancl [THEN wf_irrefl]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 207 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 208 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 209 | lemma acyclic_insert [iff]: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 210 | "acyclic(insert (y,x) r) = (acyclic r & (x,y) ~: r^*)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 211 | apply (simp add: acyclic_def trancl_insert) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 212 | apply (blast intro: rtrancl_trans) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 213 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 214 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 215 | lemma acyclic_converse [iff]: "acyclic(r^-1) = acyclic r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 216 | by (simp add: acyclic_def trancl_converse) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 217 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 218 | lemma acyclic_impl_antisym_rtrancl: "acyclic r ==> antisym(r^*)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 219 | apply (simp add: acyclic_def antisym_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 220 | apply (blast elim: rtranclE intro: rtrancl_into_trancl1 rtrancl_trancl_trancl) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 221 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 222 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 223 | (* Other direction: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 224 | acyclic = no loops | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 225 | antisym = only self loops | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 226 | Goalw [acyclic_def,antisym_def] "antisym( r^* ) ==> acyclic(r - Id) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 227 | ==> antisym( r^* ) = acyclic(r - Id)"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 228 | *) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 229 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 230 | lemma acyclic_subset: "[| acyclic s; r <= s |] ==> acyclic r" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 231 | apply (simp add: acyclic_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 232 | apply (blast intro: trancl_mono) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 233 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 234 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 235 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 236 | subsection{*Well-Founded Recursion*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 237 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 238 | text{*cut*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 239 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 240 | lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 241 | by (simp add: expand_fun_eq cut_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 242 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 243 | lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 244 | by (simp add: cut_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 245 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 246 | text{*Inductive characterization of wfrec combinator; for details see:  
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 247 | John Harrison, "Inductive definitions: automation and application"*} | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 248 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 249 | lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. (x, y) : wfrec_rel R F" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 250 | apply (simp add: adm_wf_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 251 | apply (erule_tac a=x in wf_induct) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 252 | apply (rule ex1I) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 253 | apply (rule_tac g = "%x. THE y. (x, y) : wfrec_rel R F" in wfrec_rel.wfrecI) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 254 | apply (fast dest!: theI') | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 255 | apply (erule wfrec_rel.cases, simp) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 256 | apply (erule allE, erule allE, erule allE, erule mp) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 257 | apply (fast intro: the_equality [symmetric]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 258 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 259 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 260 | lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 261 | apply (simp add: adm_wf_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 262 | apply (intro strip) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 263 | apply (rule cuts_eq [THEN iffD2, THEN subst], assumption) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 264 | apply (rule refl) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 265 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 266 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 267 | lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 268 | apply (simp add: wfrec_def) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 269 | apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 270 | apply (rule wfrec_rel.wfrecI) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 271 | apply (intro strip) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 272 | apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 273 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 274 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 275 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 276 | text{** This form avoids giant explosions in proofs.  NOTE USE OF ==*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 277 | lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 278 | apply auto | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 279 | apply (blast intro: wfrec) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 280 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 281 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 282 | |
| 17459 | 283 | subsection {* Code generator setup *}
 | 
| 284 | ||
| 285 | consts_code | |
| 17654 
38496187809d
Renamed wf_rec to wfrec in consts_code declaration.
 berghofe parents: 
17459diff
changeset | 286 |   "wfrec"   ("\<module>wfrec?")
 | 
| 17459 | 287 | attach {*
 | 
| 17654 
38496187809d
Renamed wf_rec to wfrec in consts_code declaration.
 berghofe parents: 
17459diff
changeset | 288 | fun wfrec f x = f (wfrec f) x; | 
| 17459 | 289 | *} | 
| 290 | ||
| 19602 | 291 | setup {*
 | 
| 292 |   CodegenPackage.add_appconst ("wfrec", ((3, 3), CodegenPackage.appgen_wfrec))
 | |
| 18702 | 293 | *} | 
| 294 | ||
| 19602 | 295 | code_syntax_const | 
| 18702 | 296 | wfrec | 
| 19602 | 297 |     ml ("let fun wfrec f x = f (wfrec f) x in wfrec _ _ end;")
 | 
| 298 |     haskell ("wfrec _ _ where wfrec f x = f (wfrec f) x")
 | |
| 17459 | 299 | |
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 300 | subsection{*Variants for TFL: the Recdef Package*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 301 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 302 | lemma tfl_wf_induct: "ALL R. wf R --> | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 303 | (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 304 | apply clarify | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 305 | apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 306 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 307 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 308 | lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 309 | apply clarify | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 310 | apply (rule cut_apply, assumption) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 311 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 312 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 313 | lemma tfl_wfrec: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 314 | "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 315 | apply clarify | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 316 | apply (erule wfrec) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 317 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 318 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 319 | subsection {*LEAST and wellorderings*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 320 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 321 | text{* See also @{text wf_linord_ex_has_least} and its consequences in
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 322 |  @{text Wellfounded_Relations.ML}*}
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 323 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 324 | lemma wellorder_Least_lemma [rule_format]: | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 325 | "P (k::'a::wellorder) --> P (LEAST x. P(x)) & (LEAST x. P(x)) <= k" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 326 | apply (rule_tac a = k in wf [THEN wf_induct]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 327 | apply (rule impI) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 328 | apply (rule classical) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 329 | apply (rule_tac s = x in Least_equality [THEN ssubst], auto) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 330 | apply (auto simp add: linorder_not_less [symmetric]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 331 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 332 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 333 | lemmas LeastI = wellorder_Least_lemma [THEN conjunct1, standard] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 334 | lemmas Least_le = wellorder_Least_lemma [THEN conjunct2, standard] | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 335 | |
| 15950 | 336 | -- "The following 3 lemmas are due to Brian Huffman" | 
| 337 | lemma LeastI_ex: "EX x::'a::wellorder. P x ==> P (Least P)" | |
| 338 | apply (erule exE) | |
| 339 | apply (erule LeastI) | |
| 340 | done | |
| 341 | ||
| 342 | lemma LeastI2: | |
| 343 | "[| P (a::'a::wellorder); !!x. P x ==> Q x |] ==> Q (Least P)" | |
| 344 | by (blast intro: LeastI) | |
| 345 | ||
| 346 | lemma LeastI2_ex: | |
| 347 | "[| EX a::'a::wellorder. P a; !!x. P x ==> Q x |] ==> Q (Least P)" | |
| 348 | by (blast intro: LeastI_ex) | |
| 349 | ||
| 15341 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 350 | lemma not_less_Least: "[| k < (LEAST x. P x) |] ==> ~P (k::'a::wellorder)" | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 351 | apply (simp (no_asm_use) add: linorder_not_le [symmetric]) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 352 | apply (erule contrapos_nn) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 353 | apply (erule Least_le) | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 354 | done | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 355 | |
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 356 | ML | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 357 | {*
 | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 358 | val wf_def = thm "wf_def"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 359 | val wfUNIVI = thm "wfUNIVI"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 360 | val wfI = thm "wfI"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 361 | val wf_induct = thm "wf_induct"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 362 | val wf_not_sym = thm "wf_not_sym"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 363 | val wf_asym = thm "wf_asym"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 364 | val wf_not_refl = thm "wf_not_refl"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 365 | val wf_irrefl = thm "wf_irrefl"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 366 | val wf_trancl = thm "wf_trancl"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 367 | val wf_converse_trancl = thm "wf_converse_trancl"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 368 | val wf_eq_minimal = thm "wf_eq_minimal"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 369 | val wf_subset = thm "wf_subset"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 370 | val wf_empty = thm "wf_empty"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 371 | val wf_insert = thm "wf_insert"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 372 | val wf_UN = thm "wf_UN"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 373 | val wf_Union = thm "wf_Union"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 374 | val wf_Un = thm "wf_Un"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 375 | val wf_prod_fun_image = thm "wf_prod_fun_image"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 376 | val acyclicI = thm "acyclicI"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 377 | val wf_acyclic = thm "wf_acyclic"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 378 | val acyclic_insert = thm "acyclic_insert"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 379 | val acyclic_converse = thm "acyclic_converse"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 380 | val acyclic_impl_antisym_rtrancl = thm "acyclic_impl_antisym_rtrancl"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 381 | val acyclic_subset = thm "acyclic_subset"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 382 | val cuts_eq = thm "cuts_eq"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 383 | val cut_apply = thm "cut_apply"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 384 | val wfrec_unique = thm "wfrec_unique"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 385 | val wfrec = thm "wfrec"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 386 | val def_wfrec = thm "def_wfrec"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 387 | val tfl_wf_induct = thm "tfl_wf_induct"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 388 | val tfl_cut_apply = thm "tfl_cut_apply"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 389 | val tfl_wfrec = thm "tfl_wfrec"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 390 | val LeastI = thm "LeastI"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 391 | val Least_le = thm "Least_le"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 392 | val not_less_Least = thm "not_less_Least"; | 
| 
254f6f00b60e
converted to Isar script, simplifying some results
 paulson parents: 
11451diff
changeset | 393 | *} | 
| 11137 | 394 | |
| 10213 | 395 | end |