src/HOL/HOLCF/Tutorial/New_Domain.thy
author wenzelm
Tue, 05 Jul 2016 14:20:27 +0200
changeset 63395 734723445a8c
parent 62175 8ffc4d0e652d
permissions -rw-r--r--
PIDE reports of implicit variable scope;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 40774
diff changeset
     1
(*  Title:      HOL/HOLCF/Tutorial/New_Domain.thy
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     3
*)
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>Definitional domain package\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     6
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     7
theory New_Domain
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     8
imports HOLCF
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     9
begin
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    10
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    11
text \<open>
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    12
  UPDATE: The definitional back-end is now the default mode of the domain
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    13
  package. This file should be merged with \<open>Domain_ex.thy\<close>.
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    14
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    15
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    16
text \<open>
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    17
  Provided that \<open>domain\<close> is the default sort, the \<open>new_domain\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    18
  package should work with any type definition supported by the old
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    19
  domain package.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    20
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    21
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    22
domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    23
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    24
text \<open>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    25
  The difference is that the new domain package is completely
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    26
  definitional, and does not generate any axioms.  The following type
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    27
  and constant definitions are not produced by the old domain package.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    28
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    29
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    30
thm type_definition_llist
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    31
thm llist_abs_def llist_rep_def
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    32
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    33
text \<open>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    34
  The new domain package also adds support for indirect recursion with
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    35
  user-defined datatypes.  This definition of a tree datatype uses
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    36
  indirect recursion through the lazy list type constructor.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    37
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    38
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    39
domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    40
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    41
text \<open>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    42
  For indirect-recursive definitions, the domain package is not able to
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    43
  generate a high-level induction rule.  (It produces a warning
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    44
  message instead.)  The low-level reach lemma (now proved as a
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    45
  theorem, no longer generated as an axiom) can be used to derive
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    46
  other induction rules.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    47
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    48
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    49
thm ltree.reach
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    50
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    51
text \<open>
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    52
  The definition of the take function uses map functions associated with
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    53
  each type constructor involved in the definition.  A map function
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    54
  for the lazy list type has been generated by the new domain package.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    55
\<close>
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    56
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    57
thm ltree.take_rews
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    58
thm llist_map_def
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    59
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    60
lemma ltree_induct:
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    61
  fixes P :: "'a ltree \<Rightarrow> bool"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    62
  assumes adm: "adm P"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    63
  assumes bot: "P \<bottom>"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    64
  assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    65
  assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    66
  shows "P x"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    67
proof -
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    68
  have "P (\<Squnion>i. ltree_take i\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    69
  using adm
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    70
  proof (rule admD)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    71
    fix i
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    72
    show "P (ltree_take i\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    73
    proof (induct i arbitrary: x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    74
      case (0 x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    75
      show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    76
    next
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    77
      case (Suc n x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    78
      show "P (ltree_take (Suc n)\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    79
        apply (cases x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    80
        apply (simp add: bot)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    81
        apply (simp add: Leaf)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    82
        apply (simp add: Branch Suc)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    83
        done
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    84
    qed
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    85
  qed (simp add: ltree.chain_take)
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    86
  thus ?thesis
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    87
    by (simp add: ltree.reach)
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    88
qed
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    89
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    90
end