src/HOL/Prolog/Test.thy
author haftmann
Wed, 07 Oct 2009 14:01:26 +0200
changeset 32887 85e7ab9020ba
parent 21425 c11ab38b78a7
child 34974 18b41bba42b5
permissions -rw-r--r--
added bot_boolE, top_boolI
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13208
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     1
(*  Title:    HOL/Prolog/Test.thy
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     2
    ID:       $Id$
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     3
    Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     4
*)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     5
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
header {* Basic examples *}
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     7
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
theory Test
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports HOHH
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    11
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    12
typedecl nat
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    13
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    14
typedecl 'a list
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    15
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    16
consts
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    17
  Nil   :: "'a list"                                  ("[]")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    18
  Cons  :: "'a => 'a list => 'a list"                 (infixr "#"  65)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    19
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    20
syntax
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    21
  (* list Enumeration *)
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    22
  "@list"     :: "args => 'a list"                          ("[(_)]")
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    23
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    24
translations
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    25
  "[x, xs]"     == "x#[xs]"
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    26
  "[x]"         == "x#[]"
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    27
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    28
typedecl person
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    29
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    30
consts
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    31
  append  :: "['a list, 'a list, 'a list]            => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    32
  reverse :: "['a list, 'a list]                     => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    33
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    34
  mappred :: "[('a => 'b => bool), 'a list, 'b list] => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    35
  mapfun  :: "[('a => 'b), 'a list, 'b list]         => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    36
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    37
  bob     :: person
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    38
  sue     :: person
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    39
  ned     :: person
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    40
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    41
  "23"    :: nat          ("23")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    42
  "24"    :: nat          ("24")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    43
  "25"    :: nat          ("25")
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    44
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    45
  age     :: "[person, nat]                          => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    46
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    47
  eq      :: "['a, 'a]                               => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    48
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    49
  empty   :: "['b]                                   => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    50
  add     :: "['a, 'b, 'b]                           => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    51
  remove  :: "['a, 'b, 'b]                           => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    52
  bag_appl:: "['a, 'a, 'a, 'a]                       => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    53
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    54
axioms
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    55
        append:  "append  []    xs  xs    ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    56
                  append (x#xs) ys (x#zs) :- append xs ys zs"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    57
        reverse: "reverse L1 L2 :- (!rev_aux.
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    58
                  (!L.          rev_aux  []    L  L )..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    59
                  (!X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3))
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    60
                  => rev_aux L1 L2 [])"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    61
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    62
        mappred: "mappred P  []     []    ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    63
                  mappred P (x#xs) (y#ys) :- P x y  &  mappred P xs ys"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    64
        mapfun:  "mapfun f  []     []      ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    65
                  mapfun f (x#xs) (f x#ys) :- mapfun f xs ys"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    66
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    67
        age:     "age bob 24 ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    68
                  age sue 23 ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    69
                  age ned 23"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    70
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    71
        eq:      "eq x x"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    72
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    73
(* actual definitions of empty and add is hidden -> yields abstract data type *)
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    74
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    75
        bag_appl: "bag_appl A B X Y:- (? S1 S2 S3 S4 S5.
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    76
                                empty    S1    &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    77
                                add A    S1 S2 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    78
                                add B    S2 S3 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    79
                                remove X S3 S4 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    80
                                remove Y S4 S5 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    81
                                empty    S5)"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    82
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    83
lemmas prog_Test = append reverse mappred mapfun age eq bag_appl
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    84
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    85
lemma "append ?x ?y [a,b,c,d]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    86
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    87
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    88
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    89
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    90
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    91
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    92
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    93
lemma "append [a,b] y ?L"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    94
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    95
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    96
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    97
lemma "!y. append [a,b] y (?L y)"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    98
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    99
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   100
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   101
lemma "reverse [] ?L"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   102
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   103
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   104
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   105
lemma "reverse [23] ?L"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   106
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   107
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   108
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   109
lemma "reverse [23,24,?x] ?L"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   110
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   111
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   112
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   113
lemma "reverse ?L [23,24,?x]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   114
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   115
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   116
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   117
lemma "mappred age ?x [23,24]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   118
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   119
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   120
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   121
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   122
lemma "mappred (%x y. ? z. age z y) ?x [23,24]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   123
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   124
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   125
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   126
lemma "mappred ?P [bob,sue] [24,23]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   127
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   128
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   129
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   130
lemma "mapfun f [bob,bob,sue] [?x,?y,?z]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   131
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   132
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   133
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   134
lemma "mapfun (%x. h x 25) [bob,sue] ?L"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   135
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   136
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   137
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   138
lemma "mapfun ?F [24,25] [h bob 24,h bob 25]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   139
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   140
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   141
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   142
lemma "mapfun ?F [24] [h 24 24]"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   143
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   144
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   145
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   146
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   147
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   148
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   149
lemma "True"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   150
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   151
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   152
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   153
lemma "age ?x 24 & age ?y 23"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   154
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   155
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   156
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   157
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   158
lemma "age ?x 24 | age ?x 23"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   159
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   160
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   161
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   162
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   163
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   164
lemma "? x y. age x y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   165
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   166
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   167
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   168
lemma "!x y. append [] x x"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   169
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   170
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   171
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   172
lemma "age sue 24 .. age bob 23 => age ?x ?y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   173
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   174
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   175
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   176
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   177
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   178
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   179
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   180
(*set trace_DEPTH_FIRST;*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   181
lemma "age bob 25 :- age bob 24 => age bob 25"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   182
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   183
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   184
(*reset trace_DEPTH_FIRST;*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   185
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   186
lemma "(!x. age x 25 :- age x 23) => age ?x 25 & age ?y 25"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   187
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   188
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   189
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   190
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   191
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   192
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   193
lemma "!x. ? y. eq x y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   194
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   195
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   196
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   197
lemma "? P. P & eq P ?x"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   198
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   199
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   200
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   201
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   202
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   203
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   204
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   205
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   206
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   207
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   208
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   209
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   210
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   211
lemma "? P. eq P (True & True) & P"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   212
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   213
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   214
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   215
lemma "? P. eq P op | & P k True"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   216
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   217
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   218
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   219
lemma "? P. eq P (Q => True) & P"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   220
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   221
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   222
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   223
(* P flexible: *)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   224
lemma "(!P k l. P k l :- eq P Q) => Q a b"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   225
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   226
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   227
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   228
loops:
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   229
lemma "(!P k l. P k l :- eq P (%x y. x | y)) => a | b"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   230
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   231
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   232
(* implication and disjunction in atom: *)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   233
lemma "? Q. (!p q. R(q :- p) => R(Q p q)) & Q (t | s) (s | t)"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   234
  by fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   235
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   236
(* disjunction in atom: *)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   237
lemma "(!P. g P :- (P => b | a)) => g(a | b)"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   238
  apply (tactic "step_tac HOL_cs 1")
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   239
  apply (tactic "step_tac HOL_cs 1")
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   240
  apply (tactic "step_tac HOL_cs 1")
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   241
   prefer 2
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   242
   apply fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   243
  apply fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   244
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   245
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   246
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   247
hangs:
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   248
lemma "(!P. g P :- (P => b | a)) => g(a | b)"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   249
  by fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   250
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   251
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   252
lemma "!Emp Stk.(
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   253
                       empty    (Emp::'b) .. 
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   254
         (!(X::nat) S. add    X (S::'b)         (Stk X S)) .. 
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   255
         (!(X::nat) S. remove X ((Stk X S)::'b) S))
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   256
 => bag_appl 23 24 ?X ?Y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   257
  oops
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   258
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   259
lemma "!Qu. ( 
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   260
          (!L.            empty    (Qu L L)) .. 
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   261
          (!(X::nat) L K. add    X (Qu L (X#K)) (Qu L K)) ..
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   262
          (!(X::nat) L K. remove X (Qu (X#L) K) (Qu L K)))
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   263
 => bag_appl 23 24 ?X ?Y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   264
  oops
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   265
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   266
lemma "D & (!y. E) :- (!x. True & True) :- True => D"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   267
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   268
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   269
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   270
lemma "P x .. P y => P ?X"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   271
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   272
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   273
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   274
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   275
back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   276
-> problem with DEPTH_SOLVE:
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   277
Exception- THM ("dest_state", 1, ["P x & P y --> P y"]) raised
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   278
Exception raised at run-time
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   279
*)
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
   280
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   281
end