src/HOL/Wfrec.thy
author wenzelm
Wed, 14 May 2014 12:00:18 +0200
changeset 56958 b2c2f74d1c93
parent 55210 d1e3b708d74b
child 58184 db1381d811ab
permissions -rw-r--r--
updated to polyml-5.5.2;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55210
d1e3b708d74b tuned headers;
wenzelm
parents: 55017
diff changeset
     1
(*  Title:      HOL/Wfrec.thy
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
     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
55017
2df6ad1dbd66 adapted to move of Wfrec
blanchet
parents: 55016
diff changeset
    10
imports Wellfounded
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
    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)
54482
a2874c8b3558 optimized 'bad apple' method calls
blanchet
parents: 44259
diff changeset
    51
apply (blast intro: the_equality [symmetric])
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
    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 {* 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
    78
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
    79
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
    80
 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
    81
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
    82
    "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
    83
   --{*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
    84
       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
    85
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
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
    87
     "[| 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
    88
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
    89
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
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
    91
  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
    92
  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
    93
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
    94
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
    95
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
    96
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
    97
 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
    98
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
    99
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
   100
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
end