src/ZF/InfDatatype.thy
author wenzelm
Wed, 15 Apr 2009 11:14:48 +0200
changeset 30895 bad26d8f0adf
parent 26056 6a0801279f4c
child 32960 69916a850301
permissions -rw-r--r--
updated for Isabelle2009;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     1
(*  Title:      ZF/InfDatatype.ML
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     2
    ID:         $Id$
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     4
    Copyright   1994  University of Cambridge
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     5
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     6
*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
     7
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
     8
header{*Infinite-Branching Datatype Definitions*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13269
diff changeset
     9
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents: 16417
diff changeset
    10
theory InfDatatype imports Datatype_ZF Univ Finite Cardinal_AC begin
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    11
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    12
lemmas fun_Limit_VfromE = 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    13
    Limit_VfromE [OF apply_funtype InfCard_csucc [THEN InfCard_is_Limit]]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    14
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    15
lemma fun_Vcsucc_lemma:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    16
     "[| f: D -> Vfrom(A,csucc(K));  |D| le K;  InfCard(K) |]   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    17
      ==> EX j. f: D -> Vfrom(A,j) & j < csucc(K)"
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13356
diff changeset
    18
apply (rule_tac x = "\<Union>d\<in>D. LEAST i. f`d : Vfrom (A,i) " in exI)
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    19
apply (rule conjI)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    20
apply (rule_tac [2] le_UN_Ord_lt_csucc) 
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13134
diff changeset
    21
apply (rule_tac [4] ballI, erule_tac [4] fun_Limit_VfromE, simp_all) 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    22
 prefer 2 apply (fast elim: Least_le [THEN lt_trans1] ltE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    23
apply (rule Pi_type)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    24
apply (rename_tac [2] d)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    25
apply (erule_tac [2] fun_Limit_VfromE, simp_all)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    26
apply (subgoal_tac "f`d : Vfrom (A, LEAST i. f`d : Vfrom (A,i))")
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    27
 apply (erule Vfrom_mono [OF subset_refl UN_upper, THEN subsetD])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    28
 apply assumption
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    29
apply (fast elim: LeastI ltE)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    30
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    31
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    32
lemma subset_Vcsucc:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    33
     "[| D <= Vfrom(A,csucc(K));  |D| le K;  InfCard(K) |]     
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    34
      ==> EX j. D <= Vfrom(A,j) & j < csucc(K)"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    35
by (simp add: subset_iff_id fun_Vcsucc_lemma)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    36
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    37
(*Version for arbitrary index sets*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    38
lemma fun_Vcsucc:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    39
     "[| |D| le K;  InfCard(K);  D <= Vfrom(A,csucc(K)) |] ==>  
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    40
          D -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    41
apply (safe dest!: fun_Vcsucc_lemma subset_Vcsucc)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    42
apply (rule Vfrom [THEN ssubst])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    43
apply (drule fun_is_rel)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    44
(*This level includes the function, and is below csucc(K)*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    45
apply (rule_tac a1 = "succ (succ (j Un ja))" in UN_I [THEN UnI2])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    46
apply (blast intro: ltD InfCard_csucc InfCard_is_Limit Limit_has_succ
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13134
diff changeset
    47
                    Un_least_lt) 
13134
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    48
apply (erule subset_trans [THEN PowI])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    49
apply (fast intro: Pair_in_Vfrom Vfrom_UnI1 Vfrom_UnI2)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    50
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    51
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    52
lemma fun_in_Vcsucc:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    53
     "[| f: D -> Vfrom(A, csucc(K));  |D| le K;  InfCard(K);         
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    54
         D <= Vfrom(A,csucc(K)) |]                                   
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    55
       ==> f: Vfrom(A,csucc(K))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    56
by (blast intro: fun_Vcsucc [THEN subsetD])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    57
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    58
(*Remove <= from the rule above*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    59
lemmas fun_in_Vcsucc' = fun_in_Vcsucc [OF _ _ _ subsetI]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    60
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    61
(** Version where K itself is the index set **)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    62
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    63
lemma Card_fun_Vcsucc:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    64
     "InfCard(K) ==> K -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    65
apply (frule InfCard_is_Card [THEN Card_is_Ord])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    66
apply (blast del: subsetI
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    67
	     intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    68
                   lt_csucc [THEN leI, THEN le_imp_subset, THEN subset_trans]) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    69
done
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    70
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    71
lemma Card_fun_in_Vcsucc:
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    72
     "[| f: K -> Vfrom(A, csucc(K));  InfCard(K) |] ==> f: Vfrom(A,csucc(K))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    73
by (blast intro: Card_fun_Vcsucc [THEN subsetD]) 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    74
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    75
lemma Limit_csucc: "InfCard(K) ==> Limit(csucc(K))"
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    76
by (erule InfCard_csucc [THEN InfCard_is_Limit])
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    77
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    78
lemmas Pair_in_Vcsucc = Pair_in_VLimit [OF _ _ Limit_csucc]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    79
lemmas Inl_in_Vcsucc = Inl_in_VLimit [OF _ Limit_csucc]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    80
lemmas Inr_in_Vcsucc = Inr_in_VLimit [OF _ Limit_csucc]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    81
lemmas zero_in_Vcsucc = Limit_csucc [THEN zero_in_VLimit]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    82
lemmas nat_into_Vcsucc = nat_into_VLimit [OF _ Limit_csucc]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    83
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    84
(*For handling Cardinals of the form  (nat Un |X|) *)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    85
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    86
lemmas InfCard_nat_Un_cardinal = InfCard_Un [OF InfCard_nat Card_cardinal]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    87
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    88
lemmas le_nat_Un_cardinal =
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    89
     Un_upper2_le [OF Ord_nat Card_cardinal [THEN Card_is_Ord]]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    90
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    91
lemmas UN_upper_cardinal = UN_upper [THEN subset_imp_lepoll, THEN lepoll_imp_Card_le]
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    92
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    93
(*The new version of Data_Arg.intrs, declared in Datatype.ML*)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    94
lemmas Data_Arg_intros =
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    95
       SigmaI InlI InrI
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    96
       Pair_in_univ Inl_in_univ Inr_in_univ 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    97
       zero_in_univ A_into_univ nat_into_univ UnCI
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    98
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
    99
(*For most K-branching datatypes with domain Vfrom(A, csucc(K)) *)
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   100
lemmas inf_datatype_intros =
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   101
     InfCard_nat InfCard_nat_Un_cardinal
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   102
     Pair_in_Vcsucc Inl_in_Vcsucc Inr_in_Vcsucc 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   103
     zero_in_Vcsucc A_into_Vfrom nat_into_Vcsucc
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   104
     Card_fun_in_Vcsucc fun_in_Vcsucc' UN_I 
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   105
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   106
end
bf37a3049251 converted the AC branch to Isar
paulson
parents: 516
diff changeset
   107